| author | wenzelm | 
| Thu, 27 Aug 1998 16:54:55 +0200 | |
| changeset 5393 | 7299e531d481 | 
| parent 5336 | 721bf1a13f1a | 
| child 5450 | fe9d103464a4 | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/set | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 923 | 4 | Copyright 1991 University of Cambridge | 
| 5 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1937diff
changeset | 6 | Set theory for higher-order logic. A set is simply a predicate. | 
| 923 | 7 | *) | 
| 8 | ||
| 9 | open Set; | |
| 10 | ||
| 1548 | 11 | section "Relating predicates and sets"; | 
| 12 | ||
| 3469 
61d927bd57ec
Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
 paulson parents: 
3420diff
changeset | 13 | Addsimps [Collect_mem_eq]; | 
| 
61d927bd57ec
Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
 paulson parents: 
3420diff
changeset | 14 | AddIffs [mem_Collect_eq]; | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 15 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 16 | Goal "P(a) ==> a : {x. P(x)}";
 | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 17 | by (Asm_simp_tac 1); | 
| 923 | 18 | qed "CollectI"; | 
| 19 | ||
| 5316 | 20 | Goal "a : {x. P(x)} ==> P(a)";
 | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 21 | by (Asm_full_simp_tac 1); | 
| 923 | 22 | qed "CollectD"; | 
| 23 | ||
| 5316 | 24 | val [prem] = Goal "[| !!x. (x:A) = (x:B) |] ==> A = B"; | 
| 923 | 25 | by (rtac (prem RS ext RS arg_cong RS box_equals) 1); | 
| 26 | by (rtac Collect_mem_eq 1); | |
| 27 | by (rtac Collect_mem_eq 1); | |
| 28 | qed "set_ext"; | |
| 29 | ||
| 5316 | 30 | val [prem] = Goal "[| !!x. P(x)=Q(x) |] ==> {x. P(x)} = {x. Q(x)}";
 | 
| 923 | 31 | by (rtac (prem RS ext RS arg_cong) 1); | 
| 32 | qed "Collect_cong"; | |
| 33 | ||
| 34 | val CollectE = make_elim CollectD; | |
| 35 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 36 | AddSIs [CollectI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 37 | AddSEs [CollectE]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 38 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 39 | |
| 1548 | 40 | section "Bounded quantifiers"; | 
| 923 | 41 | |
| 5316 | 42 | val prems = Goalw [Ball_def] | 
| 923 | 43 | "[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)"; | 
| 44 | by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); | |
| 45 | qed "ballI"; | |
| 46 | ||
| 5316 | 47 | Goalw [Ball_def] "[| ! x:A. P(x); x:A |] ==> P(x)"; | 
| 48 | by (Blast_tac 1); | |
| 923 | 49 | qed "bspec"; | 
| 50 | ||
| 5316 | 51 | val major::prems = Goalw [Ball_def] | 
| 923 | 52 | "[| ! x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q"; | 
| 53 | by (rtac (major RS spec RS impCE) 1); | |
| 54 | by (REPEAT (eresolve_tac prems 1)); | |
| 55 | qed "ballE"; | |
| 56 | ||
| 57 | (*Takes assumptions ! x:A.P(x) and a:A; creates assumption P(a)*) | |
| 58 | fun ball_tac i = etac ballE i THEN contr_tac (i+1); | |
| 59 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 60 | AddSIs [ballI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 61 | AddEs [ballE]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 62 | |
| 5316 | 63 | Goalw [Bex_def] "[| P(x); x:A |] ==> ? x:A. P(x)"; | 
| 64 | by (Blast_tac 1); | |
| 923 | 65 | qed "bexI"; | 
| 66 | ||
| 67 | qed_goal "bexCI" Set.thy | |
| 3842 | 68 | "[| ! x:A. ~P(x) ==> P(a); a:A |] ==> ? x:A. P(x)" | 
| 923 | 69 | (fn prems=> | 
| 70 | [ (rtac classical 1), | |
| 71 | (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]); | |
| 72 | ||
| 5316 | 73 | val major::prems = Goalw [Bex_def] | 
| 923 | 74 | "[| ? x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q"; | 
| 75 | by (rtac (major RS exE) 1); | |
| 76 | by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)); | |
| 77 | qed "bexE"; | |
| 78 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 79 | AddIs [bexI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 80 | AddSEs [bexE]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 81 | |
| 3420 | 82 | (*Trival rewrite rule*) | 
| 5069 | 83 | Goal "(! x:A. P) = ((? x. x:A) --> P)"; | 
| 4089 | 84 | by (simp_tac (simpset() addsimps [Ball_def]) 1); | 
| 3420 | 85 | qed "ball_triv"; | 
| 1816 | 86 | |
| 1882 | 87 | (*Dual form for existentials*) | 
| 5069 | 88 | Goal "(? x:A. P) = ((? x. x:A) & P)"; | 
| 4089 | 89 | by (simp_tac (simpset() addsimps [Bex_def]) 1); | 
| 3420 | 90 | qed "bex_triv"; | 
| 1882 | 91 | |
| 3420 | 92 | Addsimps [ball_triv, bex_triv]; | 
| 923 | 93 | |
| 94 | (** Congruence rules **) | |
| 95 | ||
| 5316 | 96 | val prems = Goal | 
| 923 | 97 | "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ | 
| 98 | \ (! x:A. P(x)) = (! x:B. Q(x))"; | |
| 99 | by (resolve_tac (prems RL [ssubst]) 1); | |
| 100 | by (REPEAT (ares_tac [ballI,iffI] 1 | |
| 101 | ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1)); | |
| 102 | qed "ball_cong"; | |
| 103 | ||
| 5316 | 104 | val prems = Goal | 
| 923 | 105 | "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ | 
| 106 | \ (? x:A. P(x)) = (? x:B. Q(x))"; | |
| 107 | by (resolve_tac (prems RL [ssubst]) 1); | |
| 108 | by (REPEAT (etac bexE 1 | |
| 109 | ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1)); | |
| 110 | qed "bex_cong"; | |
| 111 | ||
| 1548 | 112 | section "Subsets"; | 
| 923 | 113 | |
| 5316 | 114 | val prems = Goalw [subset_def] "(!!x. x:A ==> x:B) ==> A <= B"; | 
| 923 | 115 | by (REPEAT (ares_tac (prems @ [ballI]) 1)); | 
| 116 | qed "subsetI"; | |
| 117 | ||
| 4240 
8ba60a4cd380
Renamed "overload" to "overloaded" for sml/nj compatibility
 paulson parents: 
4231diff
changeset | 118 | Blast.overloaded ("op <=", domain_type); (*The <= relation is overloaded*)
 | 
| 4059 | 119 | |
| 120 | (*While (:) is not, its type must be kept | |
| 121 | for overloading of = to work.*) | |
| 4240 
8ba60a4cd380
Renamed "overload" to "overloaded" for sml/nj compatibility
 paulson parents: 
4231diff
changeset | 122 | Blast.overloaded ("op :", domain_type);
 | 
| 
8ba60a4cd380
Renamed "overload" to "overloaded" for sml/nj compatibility
 paulson parents: 
4231diff
changeset | 123 | seq (fn a => Blast.overloaded (a, HOLogic.dest_setT o domain_type)) | 
| 4059 | 124 | ["Ball", "Bex"]; | 
| 125 | (*need UNION, INTER also?*) | |
| 126 | ||
| 4469 | 127 | (*Image: retain the type of the set being expressed*) | 
| 5336 | 128 | Blast.overloaded ("op ``", domain_type);
 | 
| 2881 | 129 | |
| 923 | 130 | (*Rule in Modus Ponens style*) | 
| 5316 | 131 | Goalw [subset_def] "[| A <= B; c:A |] ==> c:B"; | 
| 132 | by (Blast_tac 1); | |
| 923 | 133 | qed "subsetD"; | 
| 134 | ||
| 135 | (*The same, with reversed premises for use with etac -- cf rev_mp*) | |
| 136 | qed_goal "rev_subsetD" Set.thy "[| c:A; A <= B |] ==> c:B" | |
| 137 | (fn prems=> [ (REPEAT (resolve_tac (prems@[subsetD]) 1)) ]); | |
| 138 | ||
| 1920 | 139 | (*Converts A<=B to x:A ==> x:B*) | 
| 140 | fun impOfSubs th = th RSN (2, rev_subsetD); | |
| 141 | ||
| 1841 | 142 | qed_goal "contra_subsetD" Set.thy "!!c. [| A <= B; c ~: B |] ==> c ~: A" | 
| 143 | (fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]); | |
| 144 | ||
| 145 | qed_goal "rev_contra_subsetD" Set.thy "!!c. [| c ~: B; A <= B |] ==> c ~: A" | |
| 146 | (fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]); | |
| 147 | ||
| 923 | 148 | (*Classical elimination rule*) | 
| 5316 | 149 | val major::prems = Goalw [subset_def] | 
| 923 | 150 | "[| A <= B; c~:A ==> P; c:B ==> P |] ==> P"; | 
| 151 | by (rtac (major RS ballE) 1); | |
| 152 | by (REPEAT (eresolve_tac prems 1)); | |
| 153 | qed "subsetCE"; | |
| 154 | ||
| 155 | (*Takes assumptions A<=B; c:A and creates the assumption c:B *) | |
| 156 | fun set_mp_tac i = etac subsetCE i THEN mp_tac i; | |
| 157 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 158 | AddSIs [subsetI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 159 | AddEs [subsetD, subsetCE]; | 
| 923 | 160 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 161 | qed_goal "subset_refl" Set.thy "A <= (A::'a set)" | 
| 4059 | 162 | (fn _=> [Fast_tac 1]); (*Blast_tac would try order_refl and fail*) | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 163 | |
| 5316 | 164 | Goal "[| A<=B; B<=C |] ==> A<=(C::'a set)"; | 
| 2891 | 165 | by (Blast_tac 1); | 
| 923 | 166 | qed "subset_trans"; | 
| 167 | ||
| 168 | ||
| 1548 | 169 | section "Equality"; | 
| 923 | 170 | |
| 171 | (*Anti-symmetry of the subset relation*) | |
| 5316 | 172 | Goal "[| A <= B; B <= A |] ==> A = (B::'a set)"; | 
| 5318 | 173 | by (rtac set_ext 1); | 
| 5316 | 174 | by (blast_tac (claset() addIs [subsetD]) 1); | 
| 923 | 175 | qed "subset_antisym"; | 
| 176 | val equalityI = subset_antisym; | |
| 177 | ||
| 1762 | 178 | AddSIs [equalityI]; | 
| 179 | ||
| 923 | 180 | (* Equality rules from ZF set theory -- are they appropriate here? *) | 
| 5316 | 181 | Goal "A = B ==> A<=(B::'a set)"; | 
| 182 | by (etac ssubst 1); | |
| 923 | 183 | by (rtac subset_refl 1); | 
| 184 | qed "equalityD1"; | |
| 185 | ||
| 5316 | 186 | Goal "A = B ==> B<=(A::'a set)"; | 
| 187 | by (etac ssubst 1); | |
| 923 | 188 | by (rtac subset_refl 1); | 
| 189 | qed "equalityD2"; | |
| 190 | ||
| 5316 | 191 | val prems = Goal | 
| 923 | 192 | "[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P"; | 
| 193 | by (resolve_tac prems 1); | |
| 194 | by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1)); | |
| 195 | qed "equalityE"; | |
| 196 | ||
| 5316 | 197 | val major::prems = Goal | 
| 923 | 198 | "[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P"; | 
| 199 | by (rtac (major RS equalityE) 1); | |
| 200 | by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)); | |
| 201 | qed "equalityCE"; | |
| 202 | ||
| 203 | (*Lemma for creating induction formulae -- for "pattern matching" on p | |
| 204 | To make the induction hypotheses usable, apply "spec" or "bspec" to | |
| 205 | put universal quantifiers over the free variables in p. *) | |
| 5316 | 206 | val prems = Goal | 
| 923 | 207 | "[| p:A; !!z. z:A ==> p=z --> R |] ==> R"; | 
| 208 | by (rtac mp 1); | |
| 209 | by (REPEAT (resolve_tac (refl::prems) 1)); | |
| 210 | qed "setup_induction"; | |
| 211 | ||
| 212 | ||
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 213 | section "The universal set -- UNIV"; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 214 | |
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 215 | qed_goalw "UNIV_I" Set.thy [UNIV_def] "x : UNIV" | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 216 | (fn _ => [rtac CollectI 1, rtac TrueI 1]); | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 217 | |
| 4434 | 218 | Addsimps [UNIV_I]; | 
| 219 | AddIs [UNIV_I]; (*unsafe makes it less likely to cause problems*) | |
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 220 | |
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 221 | qed_goal "subset_UNIV" Set.thy "A <= UNIV" | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 222 | (fn _ => [rtac subsetI 1, rtac UNIV_I 1]); | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 223 | |
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 224 | (** Eta-contracting these two rules (to remove P) causes them to be ignored | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 225 | because of their interaction with congruence rules. **) | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 226 | |
| 5069 | 227 | Goalw [Ball_def] "Ball UNIV P = All P"; | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 228 | by (Simp_tac 1); | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 229 | qed "ball_UNIV"; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 230 | |
| 5069 | 231 | Goalw [Bex_def] "Bex UNIV P = Ex P"; | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 232 | by (Simp_tac 1); | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 233 | qed "bex_UNIV"; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 234 | Addsimps [ball_UNIV, bex_UNIV]; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 235 | |
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 236 | |
| 2858 | 237 | section "The empty set -- {}";
 | 
| 238 | ||
| 239 | qed_goalw "empty_iff" Set.thy [empty_def] "(c : {}) = False"
 | |
| 2891 | 240 | (fn _ => [ (Blast_tac 1) ]); | 
| 2858 | 241 | |
| 242 | Addsimps [empty_iff]; | |
| 243 | ||
| 244 | qed_goal "emptyE" Set.thy "!!a. a:{} ==> P"
 | |
| 245 | (fn _ => [Full_simp_tac 1]); | |
| 246 | ||
| 247 | AddSEs [emptyE]; | |
| 248 | ||
| 249 | qed_goal "empty_subsetI" Set.thy "{} <= A"
 | |
| 2891 | 250 | (fn _ => [ (Blast_tac 1) ]); | 
| 2858 | 251 | |
| 5256 | 252 | (*One effect is to delete the ASSUMPTION {} <= A*)
 | 
| 253 | AddIffs [empty_subsetI]; | |
| 254 | ||
| 2858 | 255 | qed_goal "equals0I" Set.thy "[| !!y. y:A ==> False |] ==> A={}"
 | 
| 256 | (fn [prem]=> | |
| 4089 | 257 | [ (blast_tac (claset() addIs [prem RS FalseE]) 1) ]); | 
| 2858 | 258 | |
| 5256 | 259 | (*Use for reasoning about disjointness: A Int B = {} *)
 | 
| 260 | qed_goal "equals0E" Set.thy "!!a. [| A={};  a:A |] ==> P"
 | |
| 2891 | 261 | (fn _ => [ (Blast_tac 1) ]); | 
| 2858 | 262 | |
| 5266 | 263 | AddEs [equals0E, sym RS equals0E]; | 
| 5256 | 264 | |
| 5069 | 265 | Goalw [Ball_def] "Ball {} P = True";
 | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 266 | by (Simp_tac 1); | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 267 | qed "ball_empty"; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 268 | |
| 5069 | 269 | Goalw [Bex_def] "Bex {} P = False";
 | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 270 | by (Simp_tac 1); | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 271 | qed "bex_empty"; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 272 | Addsimps [ball_empty, bex_empty]; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 273 | |
| 5069 | 274 | Goal "UNIV ~= {}";
 | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 275 | by (blast_tac (claset() addEs [equalityE]) 1); | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 276 | qed "UNIV_not_empty"; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 277 | AddIffs [UNIV_not_empty]; | 
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 278 | |
| 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 279 | |
| 2858 | 280 | |
| 281 | section "The Powerset operator -- Pow"; | |
| 282 | ||
| 283 | qed_goalw "Pow_iff" Set.thy [Pow_def] "(A : Pow(B)) = (A <= B)" | |
| 284 | (fn _ => [ (Asm_simp_tac 1) ]); | |
| 285 | ||
| 286 | AddIffs [Pow_iff]; | |
| 287 | ||
| 288 | qed_goalw "PowI" Set.thy [Pow_def] "!!A B. A <= B ==> A : Pow(B)" | |
| 289 | (fn _ => [ (etac CollectI 1) ]); | |
| 290 | ||
| 291 | qed_goalw "PowD" Set.thy [Pow_def] "!!A B. A : Pow(B) ==> A<=B" | |
| 292 | (fn _=> [ (etac CollectD 1) ]); | |
| 293 | ||
| 294 | val Pow_bottom = empty_subsetI RS PowI;        (* {}: Pow(B) *)
 | |
| 295 | val Pow_top = subset_refl RS PowI; (* A : Pow(A) *) | |
| 296 | ||
| 297 | ||
| 1548 | 298 | section "Set complement -- Compl"; | 
| 923 | 299 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 300 | qed_goalw "Compl_iff" Set.thy [Compl_def] "(c : Compl(A)) = (c~:A)" | 
| 2891 | 301 | (fn _ => [ (Blast_tac 1) ]); | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 302 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 303 | Addsimps [Compl_iff]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 304 | |
| 5316 | 305 | val prems = Goalw [Compl_def] "[| c:A ==> False |] ==> c : Compl(A)"; | 
| 923 | 306 | by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); | 
| 307 | qed "ComplI"; | |
| 308 | ||
| 309 | (*This form, with negated conclusion, works well with the Classical prover. | |
| 310 | Negated assumptions behave like formulae on the right side of the notional | |
| 311 | turnstile...*) | |
| 5316 | 312 | Goalw [Compl_def] "c : Compl(A) ==> c~:A"; | 
| 313 | by (etac CollectD 1); | |
| 923 | 314 | qed "ComplD"; | 
| 315 | ||
| 316 | val ComplE = make_elim ComplD; | |
| 317 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 318 | AddSIs [ComplI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 319 | AddSEs [ComplE]; | 
| 1640 | 320 | |
| 923 | 321 | |
| 1548 | 322 | section "Binary union -- Un"; | 
| 923 | 323 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 324 | qed_goalw "Un_iff" Set.thy [Un_def] "(c : A Un B) = (c:A | c:B)" | 
| 2891 | 325 | (fn _ => [ Blast_tac 1 ]); | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 326 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 327 | Addsimps [Un_iff]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 328 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 329 | Goal "c:A ==> c : A Un B"; | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 330 | by (Asm_simp_tac 1); | 
| 923 | 331 | qed "UnI1"; | 
| 332 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 333 | Goal "c:B ==> c : A Un B"; | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 334 | by (Asm_simp_tac 1); | 
| 923 | 335 | qed "UnI2"; | 
| 336 | ||
| 337 | (*Classical introduction rule: no commitment to A vs B*) | |
| 338 | qed_goal "UnCI" Set.thy "(c~:B ==> c:A) ==> c : A Un B" | |
| 339 | (fn prems=> | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 340 | [ (Simp_tac 1), | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 341 | (REPEAT (ares_tac (prems@[disjCI]) 1)) ]); | 
| 923 | 342 | |
| 5316 | 343 | val major::prems = Goalw [Un_def] | 
| 923 | 344 | "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"; | 
| 345 | by (rtac (major RS CollectD RS disjE) 1); | |
| 346 | by (REPEAT (eresolve_tac prems 1)); | |
| 347 | qed "UnE"; | |
| 348 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 349 | AddSIs [UnCI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 350 | AddSEs [UnE]; | 
| 1640 | 351 | |
| 923 | 352 | |
| 1548 | 353 | section "Binary intersection -- Int"; | 
| 923 | 354 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 355 | qed_goalw "Int_iff" Set.thy [Int_def] "(c : A Int B) = (c:A & c:B)" | 
| 2891 | 356 | (fn _ => [ (Blast_tac 1) ]); | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 357 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 358 | Addsimps [Int_iff]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 359 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 360 | Goal "[| c:A; c:B |] ==> c : A Int B"; | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 361 | by (Asm_simp_tac 1); | 
| 923 | 362 | qed "IntI"; | 
| 363 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 364 | Goal "c : A Int B ==> c:A"; | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 365 | by (Asm_full_simp_tac 1); | 
| 923 | 366 | qed "IntD1"; | 
| 367 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 368 | Goal "c : A Int B ==> c:B"; | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 369 | by (Asm_full_simp_tac 1); | 
| 923 | 370 | qed "IntD2"; | 
| 371 | ||
| 5316 | 372 | val [major,minor] = Goal | 
| 923 | 373 | "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"; | 
| 374 | by (rtac minor 1); | |
| 375 | by (rtac (major RS IntD1) 1); | |
| 376 | by (rtac (major RS IntD2) 1); | |
| 377 | qed "IntE"; | |
| 378 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 379 | AddSIs [IntI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 380 | AddSEs [IntE]; | 
| 923 | 381 | |
| 1548 | 382 | section "Set difference"; | 
| 923 | 383 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 384 | qed_goalw "Diff_iff" Set.thy [set_diff_def] "(c : A-B) = (c:A & c~:B)" | 
| 2891 | 385 | (fn _ => [ (Blast_tac 1) ]); | 
| 923 | 386 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 387 | Addsimps [Diff_iff]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 388 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 389 | qed_goal "DiffI" Set.thy "!!c. [| c : A; c ~: B |] ==> c : A - B" | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 390 | (fn _=> [ Asm_simp_tac 1 ]); | 
| 923 | 391 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 392 | qed_goal "DiffD1" Set.thy "!!c. c : A - B ==> c : A" | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 393 | (fn _=> [ (Asm_full_simp_tac 1) ]); | 
| 923 | 394 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 395 | qed_goal "DiffD2" Set.thy "!!c. [| c : A - B; c : B |] ==> P" | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 396 | (fn _=> [ (Asm_full_simp_tac 1) ]); | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 397 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 398 | qed_goal "DiffE" Set.thy "[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P" | 
| 923 | 399 | (fn prems=> | 
| 400 | [ (resolve_tac prems 1), | |
| 401 | (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ]); | |
| 402 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 403 | AddSIs [DiffI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 404 | AddSEs [DiffE]; | 
| 923 | 405 | |
| 406 | ||
| 1548 | 407 | section "Augmenting a set -- insert"; | 
| 923 | 408 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 409 | qed_goalw "insert_iff" Set.thy [insert_def] "a : insert b A = (a=b | a:A)" | 
| 2891 | 410 | (fn _ => [Blast_tac 1]); | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 411 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 412 | Addsimps [insert_iff]; | 
| 923 | 413 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 414 | qed_goal "insertI1" Set.thy "a : insert a B" | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 415 | (fn _ => [Simp_tac 1]); | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 416 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 417 | qed_goal "insertI2" Set.thy "!!a. a : B ==> a : insert b B" | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 418 | (fn _=> [Asm_simp_tac 1]); | 
| 923 | 419 | |
| 420 | qed_goalw "insertE" Set.thy [insert_def] | |
| 421 | "[| a : insert b A; a=b ==> P; a:A ==> P |] ==> P" | |
| 422 | (fn major::prems=> | |
| 423 | [ (rtac (major RS UnE) 1), | |
| 424 | (REPEAT (eresolve_tac (prems @ [CollectE]) 1)) ]); | |
| 425 | ||
| 426 | (*Classical introduction rule*) | |
| 427 | qed_goal "insertCI" Set.thy "(a~:B ==> a=b) ==> a: insert b B" | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 428 | (fn prems=> | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 429 | [ (Simp_tac 1), | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 430 | (REPEAT (ares_tac (prems@[disjCI]) 1)) ]); | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 431 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 432 | AddSIs [insertCI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 433 | AddSEs [insertE]; | 
| 923 | 434 | |
| 1548 | 435 | section "Singletons, using insert"; | 
| 923 | 436 | |
| 437 | qed_goal "singletonI" Set.thy "a : {a}"
 | |
| 438 | (fn _=> [ (rtac insertI1 1) ]); | |
| 439 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 440 | Goal "b : {a} ==> b=a";
 | 
| 2891 | 441 | by (Blast_tac 1); | 
| 923 | 442 | qed "singletonD"; | 
| 443 | ||
| 1776 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 444 | bind_thm ("singletonE", make_elim singletonD);
 | 
| 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 445 | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 446 | qed_goal "singleton_iff" thy "(b : {a}) = (b=a)" 
 | 
| 2891 | 447 | (fn _ => [Blast_tac 1]); | 
| 923 | 448 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 449 | Goal "{a}={b} ==> a=b";
 | 
| 4089 | 450 | by (blast_tac (claset() addEs [equalityE]) 1); | 
| 923 | 451 | qed "singleton_inject"; | 
| 452 | ||
| 2858 | 453 | (*Redundant? But unlike insertCI, it proves the subgoal immediately!*) | 
| 454 | AddSIs [singletonI]; | |
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 455 | AddSDs [singleton_inject]; | 
| 3718 | 456 | AddSEs [singletonE]; | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 457 | |
| 5069 | 458 | Goal "{x. x=a} = {a}";
 | 
| 4423 | 459 | by (Blast_tac 1); | 
| 3582 | 460 | qed "singleton_conv"; | 
| 461 | Addsimps [singleton_conv]; | |
| 1531 | 462 | |
| 463 | ||
| 1548 | 464 | section "Unions of families -- UNION x:A. B(x) is Union(B``A)"; | 
| 923 | 465 | |
| 5069 | 466 | Goalw [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))"; | 
| 2891 | 467 | by (Blast_tac 1); | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 468 | qed "UN_iff"; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 469 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 470 | Addsimps [UN_iff]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 471 | |
| 923 | 472 | (*The order of the premises presupposes that A is rigid; b may be flexible*) | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 473 | Goal "[| a:A; b: B(a) |] ==> b: (UN x:A. B(x))"; | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4469diff
changeset | 474 | by Auto_tac; | 
| 923 | 475 | qed "UN_I"; | 
| 476 | ||
| 5316 | 477 | val major::prems = Goalw [UNION_def] | 
| 923 | 478 | "[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R"; | 
| 479 | by (rtac (major RS CollectD RS bexE) 1); | |
| 480 | by (REPEAT (ares_tac prems 1)); | |
| 481 | qed "UN_E"; | |
| 482 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 483 | AddIs [UN_I]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 484 | AddSEs [UN_E]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 485 | |
| 5316 | 486 | val prems = Goal | 
| 923 | 487 | "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ | 
| 488 | \ (UN x:A. C(x)) = (UN x:B. D(x))"; | |
| 489 | by (REPEAT (etac UN_E 1 | |
| 490 | ORELSE ares_tac ([UN_I,equalityI,subsetI] @ | |
| 1465 | 491 | (prems RL [equalityD1,equalityD2] RL [subsetD])) 1)); | 
| 923 | 492 | qed "UN_cong"; | 
| 493 | ||
| 494 | ||
| 1548 | 495 | section "Intersections of families -- INTER x:A. B(x) is Inter(B``A)"; | 
| 923 | 496 | |
| 5069 | 497 | Goalw [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))"; | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4469diff
changeset | 498 | by Auto_tac; | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 499 | qed "INT_iff"; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 500 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 501 | Addsimps [INT_iff]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 502 | |
| 5316 | 503 | val prems = Goalw [INTER_def] | 
| 923 | 504 | "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))"; | 
| 505 | by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1)); | |
| 506 | qed "INT_I"; | |
| 507 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 508 | Goal "[| b : (INT x:A. B(x)); a:A |] ==> b: B(a)"; | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4469diff
changeset | 509 | by Auto_tac; | 
| 923 | 510 | qed "INT_D"; | 
| 511 | ||
| 512 | (*"Classical" elimination -- by the Excluded Middle on a:A *) | |
| 5316 | 513 | val major::prems = Goalw [INTER_def] | 
| 923 | 514 | "[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R"; | 
| 515 | by (rtac (major RS CollectD RS ballE) 1); | |
| 516 | by (REPEAT (eresolve_tac prems 1)); | |
| 517 | qed "INT_E"; | |
| 518 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 519 | AddSIs [INT_I]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 520 | AddEs [INT_D, INT_E]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 521 | |
| 5316 | 522 | val prems = Goal | 
| 923 | 523 | "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ | 
| 524 | \ (INT x:A. C(x)) = (INT x:B. D(x))"; | |
| 525 | by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI])); | |
| 526 | by (REPEAT (dtac INT_D 1 | |
| 527 | ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1)); | |
| 528 | qed "INT_cong"; | |
| 529 | ||
| 530 | ||
| 1548 | 531 | section "Union"; | 
| 923 | 532 | |
| 5069 | 533 | Goalw [Union_def] "(A : Union(C)) = (EX X:C. A:X)"; | 
| 2891 | 534 | by (Blast_tac 1); | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 535 | qed "Union_iff"; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 536 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 537 | Addsimps [Union_iff]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 538 | |
| 923 | 539 | (*The order of the premises presupposes that C is rigid; A may be flexible*) | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 540 | Goal "[| X:C; A:X |] ==> A : Union(C)"; | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4469diff
changeset | 541 | by Auto_tac; | 
| 923 | 542 | qed "UnionI"; | 
| 543 | ||
| 5316 | 544 | val major::prems = Goalw [Union_def] | 
| 923 | 545 | "[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R"; | 
| 546 | by (rtac (major RS UN_E) 1); | |
| 547 | by (REPEAT (ares_tac prems 1)); | |
| 548 | qed "UnionE"; | |
| 549 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 550 | AddIs [UnionI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 551 | AddSEs [UnionE]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 552 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 553 | |
| 1548 | 554 | section "Inter"; | 
| 923 | 555 | |
| 5069 | 556 | Goalw [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)"; | 
| 2891 | 557 | by (Blast_tac 1); | 
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 558 | qed "Inter_iff"; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 559 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 560 | Addsimps [Inter_iff]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 561 | |
| 5316 | 562 | val prems = Goalw [Inter_def] | 
| 923 | 563 | "[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; | 
| 564 | by (REPEAT (ares_tac ([INT_I] @ prems) 1)); | |
| 565 | qed "InterI"; | |
| 566 | ||
| 567 | (*A "destruct" rule -- every X in C contains A as an element, but | |
| 568 | A:X can hold when X:C does not! This rule is analogous to "spec". *) | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 569 | Goal "[| A : Inter(C); X:C |] ==> A:X"; | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4469diff
changeset | 570 | by Auto_tac; | 
| 923 | 571 | qed "InterD"; | 
| 572 | ||
| 573 | (*"Classical" elimination rule -- does not require proving X:C *) | |
| 5316 | 574 | val major::prems = Goalw [Inter_def] | 
| 2721 | 575 | "[| A : Inter(C); X~:C ==> R; A:X ==> R |] ==> R"; | 
| 923 | 576 | by (rtac (major RS INT_E) 1); | 
| 577 | by (REPEAT (eresolve_tac prems 1)); | |
| 578 | qed "InterE"; | |
| 579 | ||
| 2499 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 580 | AddSIs [InterI]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 581 | AddEs [InterD, InterE]; | 
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 582 | |
| 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 paulson parents: 
2031diff
changeset | 583 | |
| 2912 | 584 | (*** Image of a set under a function ***) | 
| 585 | ||
| 586 | (*Frequently b does not have the syntactic form of f(x).*) | |
| 5316 | 587 | Goalw [image_def] "[| b=f(x); x:A |] ==> b : f``A"; | 
| 588 | by (Blast_tac 1); | |
| 2912 | 589 | qed "image_eqI"; | 
| 3909 | 590 | Addsimps [image_eqI]; | 
| 2912 | 591 | |
| 592 | bind_thm ("imageI", refl RS image_eqI);
 | |
| 593 | ||
| 594 | (*The eta-expansion gives variable-name preservation.*) | |
| 5316 | 595 | val major::prems = Goalw [image_def] | 
| 3842 | 596 | "[| b : (%x. f(x))``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P"; | 
| 2912 | 597 | by (rtac (major RS CollectD RS bexE) 1); | 
| 598 | by (REPEAT (ares_tac prems 1)); | |
| 599 | qed "imageE"; | |
| 600 | ||
| 601 | AddIs [image_eqI]; | |
| 602 | AddSEs [imageE]; | |
| 603 | ||
| 5069 | 604 | Goal "f``(A Un B) = f``A Un f``B"; | 
| 2935 | 605 | by (Blast_tac 1); | 
| 2912 | 606 | qed "image_Un"; | 
| 607 | ||
| 5069 | 608 | Goal "(z : f``A) = (EX x:A. z = f x)"; | 
| 3960 | 609 | by (Blast_tac 1); | 
| 610 | qed "image_iff"; | |
| 611 | ||
| 4523 | 612 | (*This rewrite rule would confuse users if made default.*) | 
| 5069 | 613 | Goal "(f``A <= B) = (ALL x:A. f(x): B)"; | 
| 4523 | 614 | by (Blast_tac 1); | 
| 615 | qed "image_subset_iff"; | |
| 616 | ||
| 617 | (*Replaces the three steps subsetI, imageE, hyp_subst_tac, but breaks too | |
| 618 | many existing proofs.*) | |
| 5316 | 619 | val prems = Goal "(!!x. x:A ==> f(x) : B) ==> f``A <= B"; | 
| 4510 | 620 | by (blast_tac (claset() addIs prems) 1); | 
| 621 | qed "image_subsetI"; | |
| 622 | ||
| 2912 | 623 | |
| 624 | (*** Range of a function -- just a translation for image! ***) | |
| 625 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 626 | Goal "b=f(x) ==> b : range(f)"; | 
| 2912 | 627 | by (EVERY1 [etac image_eqI, rtac UNIV_I]); | 
| 628 | bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI));
 | |
| 629 | ||
| 630 | bind_thm ("rangeI", UNIV_I RS imageI);
 | |
| 631 | ||
| 5316 | 632 | val [major,minor] = Goal | 
| 3842 | 633 | "[| b : range(%x. f(x)); !!x. b=f(x) ==> P |] ==> P"; | 
| 2912 | 634 | by (rtac (major RS imageE) 1); | 
| 635 | by (etac minor 1); | |
| 636 | qed "rangeE"; | |
| 637 | ||
| 1776 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 638 | |
| 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 639 | (*** Set reasoning tools ***) | 
| 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 640 | |
| 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 641 | |
| 3912 | 642 | (** Rewrite rules for boolean case-splitting: faster than | 
| 4830 | 643 | addsplits[split_if] | 
| 3912 | 644 | **) | 
| 645 | ||
| 4830 | 646 | bind_thm ("split_if_eq1", read_instantiate [("P", "%x. x = ?b")] split_if);
 | 
| 647 | bind_thm ("split_if_eq2", read_instantiate [("P", "%x. ?a = x")] split_if);
 | |
| 3912 | 648 | |
| 5237 | 649 | (*Split ifs on either side of the membership relation. | 
| 650 | Not for Addsimps -- can cause goals to blow up!*) | |
| 4830 | 651 | bind_thm ("split_if_mem1", 
 | 
| 652 |     read_instantiate_sg (sign_of Set.thy) [("P", "%x. x : ?b")] split_if);
 | |
| 653 | bind_thm ("split_if_mem2", 
 | |
| 654 |     read_instantiate_sg (sign_of Set.thy) [("P", "%x. ?a : x")] split_if);
 | |
| 3912 | 655 | |
| 4830 | 656 | val split_ifs = [if_bool_eq_conj, split_if_eq1, split_if_eq2, | 
| 657 | split_if_mem1, split_if_mem2]; | |
| 3912 | 658 | |
| 659 | ||
| 4089 | 660 | (*Each of these has ALREADY been added to simpset() above.*) | 
| 2024 
909153d8318f
Rationalized the rewriting of membership for {} and insert
 paulson parents: 
1985diff
changeset | 661 | val mem_simps = [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff, | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4135diff
changeset | 662 | mem_Collect_eq, UN_iff, Union_iff, INT_iff, Inter_iff]; | 
| 1776 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 663 | |
| 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 664 | val mksimps_pairs = ("Ball",[bspec]) :: mksimps_pairs;
 | 
| 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 665 | |
| 4089 | 666 | simpset_ref() := simpset() addcongs [ball_cong,bex_cong] | 
| 1776 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 oheimb parents: 
1762diff
changeset | 667 | setmksimps (mksimps mksimps_pairs); | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 668 | |
| 5256 | 669 | Addsimps[subset_UNIV, subset_refl]; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 670 | |
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 671 | |
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 672 | (*** < ***) | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 673 | |
| 5069 | 674 | Goalw [psubset_def] "!!A::'a set. [| A <= B; A ~= B |] ==> A<B"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 675 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 676 | qed "psubsetI"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 677 | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 678 | Goalw [psubset_def] "A < insert x B ==> (x ~: A) & A<=B | x:A & A-{x}<B";
 | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4469diff
changeset | 679 | by Auto_tac; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2935diff
changeset | 680 | qed "psubset_insertD"; | 
| 4059 | 681 | |
| 682 | bind_thm ("psubset_eq", psubset_def RS meta_eq_to_obj_eq);
 |