author | kleing |
Fri, 09 Feb 2001 16:22:30 +0100 | |
changeset 11086 | e714862ecc0a |
parent 11007 | 438f31613093 |
child 11166 | eca861fd1eb5 |
permissions | -rw-r--r-- |
9422 | 1 |
(* Title: HOL/Set.ML |
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:
1937
diff
changeset
|
6 |
Set theory for higher-order logic. A set is simply a predicate. |
923 | 7 |
*) |
8 |
||
1548 | 9 |
section "Relating predicates and sets"; |
10 |
||
3469
61d927bd57ec
Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
paulson
parents:
3420
diff
changeset
|
11 |
Addsimps [Collect_mem_eq]; |
61d927bd57ec
Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
paulson
parents:
3420
diff
changeset
|
12 |
AddIffs [mem_Collect_eq]; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
13 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
14 |
Goal "P(a) ==> a : {x. P(x)}"; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
15 |
by (Asm_simp_tac 1); |
923 | 16 |
qed "CollectI"; |
17 |
||
5316 | 18 |
Goal "a : {x. P(x)} ==> P(a)"; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
19 |
by (Asm_full_simp_tac 1); |
923 | 20 |
qed "CollectD"; |
21 |
||
9687
772ac061bd76
moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
paulson
parents:
9422
diff
changeset
|
22 |
val [prem] = Goal "(!!x. (x:A) = (x:B)) ==> A = B"; |
923 | 23 |
by (rtac (prem RS ext RS arg_cong RS box_equals) 1); |
24 |
by (rtac Collect_mem_eq 1); |
|
25 |
by (rtac Collect_mem_eq 1); |
|
26 |
qed "set_ext"; |
|
27 |
||
9687
772ac061bd76
moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
paulson
parents:
9422
diff
changeset
|
28 |
val [prem] = Goal "(!!x. P(x)=Q(x)) ==> {x. P(x)} = {x. Q(x)}"; |
923 | 29 |
by (rtac (prem RS ext RS arg_cong) 1); |
30 |
qed "Collect_cong"; |
|
31 |
||
9108 | 32 |
bind_thm ("CollectE", make_elim CollectD); |
923 | 33 |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
34 |
AddSIs [CollectI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
35 |
AddSEs [CollectE]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
36 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
37 |
|
1548 | 38 |
section "Bounded quantifiers"; |
923 | 39 |
|
5316 | 40 |
val prems = Goalw [Ball_def] |
9041 | 41 |
"[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)"; |
923 | 42 |
by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); |
43 |
qed "ballI"; |
|
44 |
||
8839 | 45 |
bind_thms ("strip", [impI, allI, ballI]); |
46 |
||
9041 | 47 |
Goalw [Ball_def] "[| ALL x:A. P(x); x:A |] ==> P(x)"; |
5316 | 48 |
by (Blast_tac 1); |
923 | 49 |
qed "bspec"; |
50 |
||
5316 | 51 |
val major::prems = Goalw [Ball_def] |
9041 | 52 |
"[| ALL x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q"; |
923 | 53 |
by (rtac (major RS spec RS impCE) 1); |
54 |
by (REPEAT (eresolve_tac prems 1)); |
|
55 |
qed "ballE"; |
|
56 |
||
9041 | 57 |
(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*) |
923 | 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:
2031
diff
changeset
|
60 |
AddSIs [ballI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
61 |
AddEs [ballE]; |
7441 | 62 |
AddXDs [bspec]; |
5521 | 63 |
(* gives better instantiation for bound: *) |
64 |
claset_ref() := claset() addWrapper ("bspec", fn tac2 => |
|
65 |
(dtac bspec THEN' atac) APPEND' tac2); |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
66 |
|
6006 | 67 |
(*Normally the best argument order: P(x) constrains the choice of x:A*) |
9041 | 68 |
Goalw [Bex_def] "[| P(x); x:A |] ==> EX x:A. P(x)"; |
5316 | 69 |
by (Blast_tac 1); |
923 | 70 |
qed "bexI"; |
71 |
||
6006 | 72 |
(*The best argument order when there is only one x:A*) |
9041 | 73 |
Goalw [Bex_def] "[| x:A; P(x) |] ==> EX x:A. P(x)"; |
6006 | 74 |
by (Blast_tac 1); |
75 |
qed "rev_bexI"; |
|
76 |
||
7031 | 77 |
val prems = Goal |
9041 | 78 |
"[| ALL x:A. ~P(x) ==> P(a); a:A |] ==> EX x:A. P(x)"; |
7007 | 79 |
by (rtac classical 1); |
80 |
by (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ; |
|
81 |
qed "bexCI"; |
|
923 | 82 |
|
5316 | 83 |
val major::prems = Goalw [Bex_def] |
9041 | 84 |
"[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q"; |
923 | 85 |
by (rtac (major RS exE) 1); |
86 |
by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)); |
|
87 |
qed "bexE"; |
|
88 |
||
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
89 |
AddIs [bexI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
90 |
AddSEs [bexE]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
91 |
|
3420 | 92 |
(*Trival rewrite rule*) |
9041 | 93 |
Goal "(ALL x:A. P) = ((EX x. x:A) --> P)"; |
4089 | 94 |
by (simp_tac (simpset() addsimps [Ball_def]) 1); |
3420 | 95 |
qed "ball_triv"; |
1816 | 96 |
|
1882 | 97 |
(*Dual form for existentials*) |
9041 | 98 |
Goal "(EX x:A. P) = ((EX x. x:A) & P)"; |
4089 | 99 |
by (simp_tac (simpset() addsimps [Bex_def]) 1); |
3420 | 100 |
qed "bex_triv"; |
1882 | 101 |
|
3420 | 102 |
Addsimps [ball_triv, bex_triv]; |
923 | 103 |
|
104 |
(** Congruence rules **) |
|
105 |
||
6291 | 106 |
val prems = Goalw [Ball_def] |
923 | 107 |
"[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ |
9041 | 108 |
\ (ALL x:A. P(x)) = (ALL x:B. Q(x))"; |
6291 | 109 |
by (asm_simp_tac (simpset() addsimps prems) 1); |
923 | 110 |
qed "ball_cong"; |
111 |
||
6291 | 112 |
val prems = Goalw [Bex_def] |
923 | 113 |
"[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ |
9041 | 114 |
\ (EX x:A. P(x)) = (EX x:B. Q(x))"; |
6291 | 115 |
by (asm_simp_tac (simpset() addcongs [conj_cong] addsimps prems) 1); |
923 | 116 |
qed "bex_cong"; |
117 |
||
6291 | 118 |
Addcongs [ball_cong,bex_cong]; |
119 |
||
1548 | 120 |
section "Subsets"; |
923 | 121 |
|
5316 | 122 |
val prems = Goalw [subset_def] "(!!x. x:A ==> x:B) ==> A <= B"; |
923 | 123 |
by (REPEAT (ares_tac (prems @ [ballI]) 1)); |
124 |
qed "subsetI"; |
|
125 |
||
5649 | 126 |
(*Map the type ('a set => anything) to just 'a. |
127 |
For overloading constants whose first argument has type "'a set" *) |
|
128 |
fun overload_1st_set s = Blast.overloaded (s, HOLogic.dest_setT o domain_type); |
|
129 |
||
4059 | 130 |
(*While (:) is not, its type must be kept |
131 |
for overloading of = to work.*) |
|
4240
8ba60a4cd380
Renamed "overload" to "overloaded" for sml/nj compatibility
paulson
parents:
4231
diff
changeset
|
132 |
Blast.overloaded ("op :", domain_type); |
5649 | 133 |
|
134 |
overload_1st_set "Ball"; (*need UNION, INTER also?*) |
|
135 |
overload_1st_set "Bex"; |
|
4059 | 136 |
|
4469 | 137 |
(*Image: retain the type of the set being expressed*) |
8005 | 138 |
Blast.overloaded ("image", domain_type); |
2881 | 139 |
|
923 | 140 |
(*Rule in Modus Ponens style*) |
5316 | 141 |
Goalw [subset_def] "[| A <= B; c:A |] ==> c:B"; |
142 |
by (Blast_tac 1); |
|
923 | 143 |
qed "subsetD"; |
7658 | 144 |
AddXIs [subsetD]; |
923 | 145 |
|
146 |
(*The same, with reversed premises for use with etac -- cf rev_mp*) |
|
7007 | 147 |
Goal "[| c:A; A <= B |] ==> c:B"; |
148 |
by (REPEAT (ares_tac [subsetD] 1)) ; |
|
149 |
qed "rev_subsetD"; |
|
7658 | 150 |
AddXIs [rev_subsetD]; |
923 | 151 |
|
1920 | 152 |
(*Converts A<=B to x:A ==> x:B*) |
153 |
fun impOfSubs th = th RSN (2, rev_subsetD); |
|
154 |
||
923 | 155 |
(*Classical elimination rule*) |
5316 | 156 |
val major::prems = Goalw [subset_def] |
923 | 157 |
"[| A <= B; c~:A ==> P; c:B ==> P |] ==> P"; |
158 |
by (rtac (major RS ballE) 1); |
|
159 |
by (REPEAT (eresolve_tac prems 1)); |
|
160 |
qed "subsetCE"; |
|
161 |
||
162 |
(*Takes assumptions A<=B; c:A and creates the assumption c:B *) |
|
163 |
fun set_mp_tac i = etac subsetCE i THEN mp_tac i; |
|
164 |
||
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
165 |
AddSIs [subsetI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
166 |
AddEs [subsetD, subsetCE]; |
923 | 167 |
|
10233 | 168 |
Goal "[| A <= B; c ~: B |] ==> c ~: A"; |
169 |
by (Blast_tac 1); |
|
170 |
qed "contra_subsetD"; |
|
171 |
||
7007 | 172 |
Goal "A <= (A::'a set)"; |
173 |
by (Fast_tac 1); |
|
10233 | 174 |
qed "subset_refl"; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
175 |
|
5316 | 176 |
Goal "[| A<=B; B<=C |] ==> A<=(C::'a set)"; |
2891 | 177 |
by (Blast_tac 1); |
923 | 178 |
qed "subset_trans"; |
179 |
||
180 |
||
1548 | 181 |
section "Equality"; |
923 | 182 |
|
183 |
(*Anti-symmetry of the subset relation*) |
|
5316 | 184 |
Goal "[| A <= B; B <= A |] ==> A = (B::'a set)"; |
5318 | 185 |
by (rtac set_ext 1); |
5316 | 186 |
by (blast_tac (claset() addIs [subsetD]) 1); |
923 | 187 |
qed "subset_antisym"; |
9108 | 188 |
bind_thm ("equalityI", subset_antisym); |
923 | 189 |
|
1762 | 190 |
AddSIs [equalityI]; |
191 |
||
923 | 192 |
(* Equality rules from ZF set theory -- are they appropriate here? *) |
5316 | 193 |
Goal "A = B ==> A<=(B::'a set)"; |
194 |
by (etac ssubst 1); |
|
923 | 195 |
by (rtac subset_refl 1); |
196 |
qed "equalityD1"; |
|
197 |
||
5316 | 198 |
Goal "A = B ==> B<=(A::'a set)"; |
199 |
by (etac ssubst 1); |
|
923 | 200 |
by (rtac subset_refl 1); |
201 |
qed "equalityD2"; |
|
202 |
||
9338 | 203 |
(*Be careful when adding this to the claset as subset_empty is in the simpset: |
204 |
A={} goes to {}<=A and A<={} and then back to A={} !*) |
|
5316 | 205 |
val prems = Goal |
923 | 206 |
"[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P"; |
207 |
by (resolve_tac prems 1); |
|
208 |
by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1)); |
|
209 |
qed "equalityE"; |
|
210 |
||
5316 | 211 |
val major::prems = Goal |
923 | 212 |
"[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P"; |
213 |
by (rtac (major RS equalityE) 1); |
|
214 |
by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)); |
|
215 |
qed "equalityCE"; |
|
216 |
||
9186
7b2f4e6538b4
now uses equalityCE, which usually is more efficent than equalityE
paulson
parents:
9161
diff
changeset
|
217 |
AddEs [equalityCE]; |
7b2f4e6538b4
now uses equalityCE, which usually is more efficent than equalityE
paulson
parents:
9161
diff
changeset
|
218 |
|
923 | 219 |
(*Lemma for creating induction formulae -- for "pattern matching" on p |
220 |
To make the induction hypotheses usable, apply "spec" or "bspec" to |
|
221 |
put universal quantifiers over the free variables in p. *) |
|
5316 | 222 |
val prems = Goal |
923 | 223 |
"[| p:A; !!z. z:A ==> p=z --> R |] ==> R"; |
224 |
by (rtac mp 1); |
|
225 |
by (REPEAT (resolve_tac (refl::prems) 1)); |
|
226 |
qed "setup_induction"; |
|
227 |
||
8053 | 228 |
Goal "A = B ==> (x : A) = (x : B)"; |
229 |
by (Asm_simp_tac 1); |
|
230 |
qed "eqset_imp_iff"; |
|
231 |
||
923 | 232 |
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
233 |
section "The universal set -- UNIV"; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
234 |
|
7031 | 235 |
Goalw [UNIV_def] "x : UNIV"; |
236 |
by (rtac CollectI 1); |
|
237 |
by (rtac TrueI 1); |
|
238 |
qed "UNIV_I"; |
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
239 |
|
4434 | 240 |
Addsimps [UNIV_I]; |
241 |
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:
4135
diff
changeset
|
242 |
|
10482 | 243 |
Goal "EX x. x : UNIV"; |
244 |
by (Simp_tac 1); |
|
245 |
qed "UNIV_witness"; |
|
246 |
AddXIs [UNIV_witness]; |
|
247 |
||
7031 | 248 |
Goal "A <= UNIV"; |
249 |
by (rtac subsetI 1); |
|
250 |
by (rtac UNIV_I 1); |
|
251 |
qed "subset_UNIV"; |
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
252 |
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
253 |
(** 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:
4135
diff
changeset
|
254 |
because of their interaction with congruence rules. **) |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
255 |
|
5069 | 256 |
Goalw [Ball_def] "Ball UNIV P = All P"; |
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
257 |
by (Simp_tac 1); |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
258 |
qed "ball_UNIV"; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
259 |
|
5069 | 260 |
Goalw [Bex_def] "Bex UNIV P = Ex P"; |
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
261 |
by (Simp_tac 1); |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
262 |
qed "bex_UNIV"; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
263 |
Addsimps [ball_UNIV, bex_UNIV]; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
264 |
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
265 |
|
2858 | 266 |
section "The empty set -- {}"; |
267 |
||
7007 | 268 |
Goalw [empty_def] "(c : {}) = False"; |
269 |
by (Blast_tac 1) ; |
|
270 |
qed "empty_iff"; |
|
2858 | 271 |
|
272 |
Addsimps [empty_iff]; |
|
273 |
||
7007 | 274 |
Goal "a:{} ==> P"; |
275 |
by (Full_simp_tac 1); |
|
276 |
qed "emptyE"; |
|
2858 | 277 |
|
278 |
AddSEs [emptyE]; |
|
279 |
||
7007 | 280 |
Goal "{} <= A"; |
281 |
by (Blast_tac 1) ; |
|
282 |
qed "empty_subsetI"; |
|
2858 | 283 |
|
5256 | 284 |
(*One effect is to delete the ASSUMPTION {} <= A*) |
285 |
AddIffs [empty_subsetI]; |
|
286 |
||
7031 | 287 |
val [prem]= Goal "[| !!y. y:A ==> False |] ==> A={}"; |
7007 | 288 |
by (blast_tac (claset() addIs [prem RS FalseE]) 1) ; |
289 |
qed "equals0I"; |
|
2858 | 290 |
|
5256 | 291 |
(*Use for reasoning about disjointness: A Int B = {} *) |
7007 | 292 |
Goal "A={} ==> a ~: A"; |
293 |
by (Blast_tac 1) ; |
|
294 |
qed "equals0D"; |
|
2858 | 295 |
|
5069 | 296 |
Goalw [Ball_def] "Ball {} P = True"; |
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
297 |
by (Simp_tac 1); |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
298 |
qed "ball_empty"; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
299 |
|
5069 | 300 |
Goalw [Bex_def] "Bex {} P = False"; |
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
301 |
by (Simp_tac 1); |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
302 |
qed "bex_empty"; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
303 |
Addsimps [ball_empty, bex_empty]; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
304 |
|
5069 | 305 |
Goal "UNIV ~= {}"; |
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
306 |
by (blast_tac (claset() addEs [equalityE]) 1); |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
307 |
qed "UNIV_not_empty"; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
308 |
AddIffs [UNIV_not_empty]; |
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
309 |
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4135
diff
changeset
|
310 |
|
2858 | 311 |
|
312 |
section "The Powerset operator -- Pow"; |
|
313 |
||
7007 | 314 |
Goalw [Pow_def] "(A : Pow(B)) = (A <= B)"; |
315 |
by (Asm_simp_tac 1); |
|
316 |
qed "Pow_iff"; |
|
2858 | 317 |
|
318 |
AddIffs [Pow_iff]; |
|
319 |
||
7031 | 320 |
Goalw [Pow_def] "A <= B ==> A : Pow(B)"; |
7007 | 321 |
by (etac CollectI 1); |
322 |
qed "PowI"; |
|
2858 | 323 |
|
7031 | 324 |
Goalw [Pow_def] "A : Pow(B) ==> A<=B"; |
7007 | 325 |
by (etac CollectD 1); |
326 |
qed "PowD"; |
|
327 |
||
2858 | 328 |
|
9108 | 329 |
bind_thm ("Pow_bottom", empty_subsetI RS PowI); (* {}: Pow(B) *) |
330 |
bind_thm ("Pow_top", subset_refl RS PowI); (* A : Pow(A) *) |
|
2858 | 331 |
|
332 |
||
5931 | 333 |
section "Set complement"; |
923 | 334 |
|
7031 | 335 |
Goalw [Compl_def] "(c : -A) = (c~:A)"; |
336 |
by (Blast_tac 1); |
|
337 |
qed "Compl_iff"; |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
338 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
339 |
Addsimps [Compl_iff]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
340 |
|
5490 | 341 |
val prems = Goalw [Compl_def] "[| c:A ==> False |] ==> c : -A"; |
923 | 342 |
by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); |
343 |
qed "ComplI"; |
|
344 |
||
345 |
(*This form, with negated conclusion, works well with the Classical prover. |
|
346 |
Negated assumptions behave like formulae on the right side of the notional |
|
347 |
turnstile...*) |
|
5490 | 348 |
Goalw [Compl_def] "c : -A ==> c~:A"; |
5316 | 349 |
by (etac CollectD 1); |
923 | 350 |
qed "ComplD"; |
351 |
||
9108 | 352 |
bind_thm ("ComplE", make_elim ComplD); |
923 | 353 |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
354 |
AddSIs [ComplI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
355 |
AddSEs [ComplE]; |
1640 | 356 |
|
923 | 357 |
|
1548 | 358 |
section "Binary union -- Un"; |
923 | 359 |
|
7031 | 360 |
Goalw [Un_def] "(c : A Un B) = (c:A | c:B)"; |
361 |
by (Blast_tac 1); |
|
362 |
qed "Un_iff"; |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
363 |
Addsimps [Un_iff]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
364 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
365 |
Goal "c:A ==> c : A Un B"; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
366 |
by (Asm_simp_tac 1); |
923 | 367 |
qed "UnI1"; |
368 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
369 |
Goal "c:B ==> c : A Un B"; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
370 |
by (Asm_simp_tac 1); |
923 | 371 |
qed "UnI2"; |
372 |
||
9378 | 373 |
AddXIs [UnI1, UnI2]; |
374 |
||
375 |
||
923 | 376 |
(*Classical introduction rule: no commitment to A vs B*) |
7007 | 377 |
|
7031 | 378 |
val prems = Goal "(c~:B ==> c:A) ==> c : A Un B"; |
7007 | 379 |
by (Simp_tac 1); |
380 |
by (REPEAT (ares_tac (prems@[disjCI]) 1)) ; |
|
381 |
qed "UnCI"; |
|
923 | 382 |
|
5316 | 383 |
val major::prems = Goalw [Un_def] |
923 | 384 |
"[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"; |
385 |
by (rtac (major RS CollectD RS disjE) 1); |
|
386 |
by (REPEAT (eresolve_tac prems 1)); |
|
387 |
qed "UnE"; |
|
388 |
||
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
389 |
AddSIs [UnCI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
390 |
AddSEs [UnE]; |
1640 | 391 |
|
923 | 392 |
|
1548 | 393 |
section "Binary intersection -- Int"; |
923 | 394 |
|
7031 | 395 |
Goalw [Int_def] "(c : A Int B) = (c:A & c:B)"; |
396 |
by (Blast_tac 1); |
|
397 |
qed "Int_iff"; |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
398 |
Addsimps [Int_iff]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
399 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
400 |
Goal "[| c:A; c:B |] ==> c : A Int B"; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
401 |
by (Asm_simp_tac 1); |
923 | 402 |
qed "IntI"; |
403 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
404 |
Goal "c : A Int B ==> c:A"; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
405 |
by (Asm_full_simp_tac 1); |
923 | 406 |
qed "IntD1"; |
407 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
408 |
Goal "c : A Int B ==> c:B"; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
409 |
by (Asm_full_simp_tac 1); |
923 | 410 |
qed "IntD2"; |
411 |
||
5316 | 412 |
val [major,minor] = Goal |
923 | 413 |
"[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"; |
414 |
by (rtac minor 1); |
|
415 |
by (rtac (major RS IntD1) 1); |
|
416 |
by (rtac (major RS IntD2) 1); |
|
417 |
qed "IntE"; |
|
418 |
||
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
419 |
AddSIs [IntI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
420 |
AddSEs [IntE]; |
923 | 421 |
|
1548 | 422 |
section "Set difference"; |
923 | 423 |
|
7031 | 424 |
Goalw [set_diff_def] "(c : A-B) = (c:A & c~:B)"; |
425 |
by (Blast_tac 1); |
|
426 |
qed "Diff_iff"; |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
427 |
Addsimps [Diff_iff]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
428 |
|
7007 | 429 |
Goal "[| c : A; c ~: B |] ==> c : A - B"; |
430 |
by (Asm_simp_tac 1) ; |
|
431 |
qed "DiffI"; |
|
923 | 432 |
|
7007 | 433 |
Goal "c : A - B ==> c : A"; |
434 |
by (Asm_full_simp_tac 1) ; |
|
435 |
qed "DiffD1"; |
|
923 | 436 |
|
7007 | 437 |
Goal "[| c : A - B; c : B |] ==> P"; |
438 |
by (Asm_full_simp_tac 1) ; |
|
439 |
qed "DiffD2"; |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
440 |
|
7031 | 441 |
val prems = Goal "[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P"; |
7007 | 442 |
by (resolve_tac prems 1); |
443 |
by (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ; |
|
444 |
qed "DiffE"; |
|
923 | 445 |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
446 |
AddSIs [DiffI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
447 |
AddSEs [DiffE]; |
923 | 448 |
|
449 |
||
1548 | 450 |
section "Augmenting a set -- insert"; |
923 | 451 |
|
7031 | 452 |
Goalw [insert_def] "a : insert b A = (a=b | a:A)"; |
453 |
by (Blast_tac 1); |
|
454 |
qed "insert_iff"; |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
455 |
Addsimps [insert_iff]; |
923 | 456 |
|
7031 | 457 |
Goal "a : insert a B"; |
7007 | 458 |
by (Simp_tac 1); |
459 |
qed "insertI1"; |
|
923 | 460 |
|
7007 | 461 |
Goal "!!a. a : B ==> a : insert b B"; |
462 |
by (Asm_simp_tac 1); |
|
463 |
qed "insertI2"; |
|
464 |
||
465 |
val major::prems = Goalw [insert_def] |
|
466 |
"[| a : insert b A; a=b ==> P; a:A ==> P |] ==> P"; |
|
467 |
by (rtac (major RS UnE) 1); |
|
468 |
by (REPEAT (eresolve_tac (prems @ [CollectE]) 1)); |
|
469 |
qed "insertE"; |
|
923 | 470 |
|
471 |
(*Classical introduction rule*) |
|
7031 | 472 |
val prems = Goal "(a~:B ==> a=b) ==> a: insert b B"; |
7007 | 473 |
by (Simp_tac 1); |
474 |
by (REPEAT (ares_tac (prems@[disjCI]) 1)) ; |
|
475 |
qed "insertCI"; |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
476 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
477 |
AddSIs [insertCI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
478 |
AddSEs [insertE]; |
923 | 479 |
|
9088
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents:
9075
diff
changeset
|
480 |
Goal "(A <= insert x B) = (if x:A then A-{x} <= B else A<=B)"; |
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents:
9075
diff
changeset
|
481 |
by Auto_tac; |
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents:
9075
diff
changeset
|
482 |
qed "subset_insert_iff"; |
7496
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents:
7441
diff
changeset
|
483 |
|
1548 | 484 |
section "Singletons, using insert"; |
923 | 485 |
|
7007 | 486 |
Goal "a : {a}"; |
487 |
by (rtac insertI1 1) ; |
|
488 |
qed "singletonI"; |
|
923 | 489 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
490 |
Goal "b : {a} ==> b=a"; |
2891 | 491 |
by (Blast_tac 1); |
923 | 492 |
qed "singletonD"; |
493 |
||
1776
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
494 |
bind_thm ("singletonE", make_elim singletonD); |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
495 |
|
7007 | 496 |
Goal "(b : {a}) = (b=a)"; |
497 |
by (Blast_tac 1); |
|
498 |
qed "singleton_iff"; |
|
923 | 499 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
500 |
Goal "{a}={b} ==> a=b"; |
4089 | 501 |
by (blast_tac (claset() addEs [equalityE]) 1); |
923 | 502 |
qed "singleton_inject"; |
503 |
||
2858 | 504 |
(*Redundant? But unlike insertCI, it proves the subgoal immediately!*) |
505 |
AddSIs [singletonI]; |
|
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
506 |
AddSDs [singleton_inject]; |
3718 | 507 |
AddSEs [singletonE]; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
508 |
|
7969 | 509 |
Goal "{b} = insert a A = (a = b & A <= {b})"; |
8326
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents:
8053
diff
changeset
|
510 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
7496
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents:
7441
diff
changeset
|
511 |
qed "singleton_insert_inj_eq"; |
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents:
7441
diff
changeset
|
512 |
|
8326
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents:
8053
diff
changeset
|
513 |
Goal "(insert a A = {b}) = (a = b & A <= {b})"; |
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents:
8053
diff
changeset
|
514 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
7969 | 515 |
qed "singleton_insert_inj_eq'"; |
516 |
||
8326
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents:
8053
diff
changeset
|
517 |
AddIffs [singleton_insert_inj_eq, singleton_insert_inj_eq']; |
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
paulson
parents:
8053
diff
changeset
|
518 |
|
7496
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents:
7441
diff
changeset
|
519 |
Goal "A <= {x} ==> A={} | A = {x}"; |
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents:
7441
diff
changeset
|
520 |
by (Fast_tac 1); |
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents:
7441
diff
changeset
|
521 |
qed "subset_singletonD"; |
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
oheimb
parents:
7441
diff
changeset
|
522 |
|
5069 | 523 |
Goal "{x. x=a} = {a}"; |
4423 | 524 |
by (Blast_tac 1); |
3582 | 525 |
qed "singleton_conv"; |
526 |
Addsimps [singleton_conv]; |
|
1531 | 527 |
|
5600 | 528 |
Goal "{x. a=x} = {a}"; |
6301 | 529 |
by (Blast_tac 1); |
5600 | 530 |
qed "singleton_conv2"; |
531 |
Addsimps [singleton_conv2]; |
|
532 |
||
11007 | 533 |
Goal "[| A - {x} <= B; x : A |] ==> A <= insert x B"; |
534 |
by(Blast_tac 1); |
|
535 |
qed "diff_single_insert"; |
|
536 |
||
1531 | 537 |
|
10832 | 538 |
section "Unions of families -- UNION x:A. B(x) is Union(B`A)"; |
923 | 539 |
|
5069 | 540 |
Goalw [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))"; |
2891 | 541 |
by (Blast_tac 1); |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
542 |
qed "UN_iff"; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
543 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
544 |
Addsimps [UN_iff]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
545 |
|
923 | 546 |
(*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:
5069
diff
changeset
|
547 |
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:
4469
diff
changeset
|
548 |
by Auto_tac; |
923 | 549 |
qed "UN_I"; |
550 |
||
5316 | 551 |
val major::prems = Goalw [UNION_def] |
923 | 552 |
"[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R"; |
553 |
by (rtac (major RS CollectD RS bexE) 1); |
|
554 |
by (REPEAT (ares_tac prems 1)); |
|
555 |
qed "UN_E"; |
|
556 |
||
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
557 |
AddIs [UN_I]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
558 |
AddSEs [UN_E]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
559 |
|
6291 | 560 |
val prems = Goalw [UNION_def] |
923 | 561 |
"[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ |
562 |
\ (UN x:A. C(x)) = (UN x:B. D(x))"; |
|
6291 | 563 |
by (asm_simp_tac (simpset() addsimps prems) 1); |
923 | 564 |
qed "UN_cong"; |
9687
772ac061bd76
moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
paulson
parents:
9422
diff
changeset
|
565 |
Addcongs [UN_cong]; |
923 | 566 |
|
567 |
||
10832 | 568 |
section "Intersections of families -- INTER x:A. B(x) is Inter(B`A)"; |
923 | 569 |
|
5069 | 570 |
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:
4469
diff
changeset
|
571 |
by Auto_tac; |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
572 |
qed "INT_iff"; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
573 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
574 |
Addsimps [INT_iff]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
575 |
|
5316 | 576 |
val prems = Goalw [INTER_def] |
923 | 577 |
"(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))"; |
578 |
by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1)); |
|
579 |
qed "INT_I"; |
|
580 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
581 |
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:
4469
diff
changeset
|
582 |
by Auto_tac; |
923 | 583 |
qed "INT_D"; |
584 |
||
585 |
(*"Classical" elimination -- by the Excluded Middle on a:A *) |
|
5316 | 586 |
val major::prems = Goalw [INTER_def] |
923 | 587 |
"[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R"; |
588 |
by (rtac (major RS CollectD RS ballE) 1); |
|
589 |
by (REPEAT (eresolve_tac prems 1)); |
|
590 |
qed "INT_E"; |
|
591 |
||
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
592 |
AddSIs [INT_I]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
593 |
AddEs [INT_D, INT_E]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
594 |
|
6291 | 595 |
val prems = Goalw [INTER_def] |
923 | 596 |
"[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ |
597 |
\ (INT x:A. C(x)) = (INT x:B. D(x))"; |
|
6291 | 598 |
by (asm_simp_tac (simpset() addsimps prems) 1); |
923 | 599 |
qed "INT_cong"; |
9687
772ac061bd76
moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
paulson
parents:
9422
diff
changeset
|
600 |
Addcongs [INT_cong]; |
923 | 601 |
|
602 |
||
1548 | 603 |
section "Union"; |
923 | 604 |
|
5069 | 605 |
Goalw [Union_def] "(A : Union(C)) = (EX X:C. A:X)"; |
2891 | 606 |
by (Blast_tac 1); |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
607 |
qed "Union_iff"; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
608 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
609 |
Addsimps [Union_iff]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
610 |
|
923 | 611 |
(*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:
5069
diff
changeset
|
612 |
Goal "[| X:C; A:X |] ==> A : Union(C)"; |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4469
diff
changeset
|
613 |
by Auto_tac; |
923 | 614 |
qed "UnionI"; |
615 |
||
5316 | 616 |
val major::prems = Goalw [Union_def] |
923 | 617 |
"[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R"; |
618 |
by (rtac (major RS UN_E) 1); |
|
619 |
by (REPEAT (ares_tac prems 1)); |
|
620 |
qed "UnionE"; |
|
621 |
||
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
622 |
AddIs [UnionI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
623 |
AddSEs [UnionE]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
624 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
625 |
|
1548 | 626 |
section "Inter"; |
923 | 627 |
|
5069 | 628 |
Goalw [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)"; |
2891 | 629 |
by (Blast_tac 1); |
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
630 |
qed "Inter_iff"; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
631 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
632 |
Addsimps [Inter_iff]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
633 |
|
5316 | 634 |
val prems = Goalw [Inter_def] |
923 | 635 |
"[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; |
636 |
by (REPEAT (ares_tac ([INT_I] @ prems) 1)); |
|
637 |
qed "InterI"; |
|
638 |
||
639 |
(*A "destruct" rule -- every X in C contains A as an element, but |
|
640 |
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:
5069
diff
changeset
|
641 |
Goal "[| A : Inter(C); X:C |] ==> A:X"; |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4469
diff
changeset
|
642 |
by Auto_tac; |
923 | 643 |
qed "InterD"; |
644 |
||
645 |
(*"Classical" elimination rule -- does not require proving X:C *) |
|
5316 | 646 |
val major::prems = Goalw [Inter_def] |
2721 | 647 |
"[| A : Inter(C); X~:C ==> R; A:X ==> R |] ==> R"; |
923 | 648 |
by (rtac (major RS INT_E) 1); |
649 |
by (REPEAT (eresolve_tac prems 1)); |
|
650 |
qed "InterE"; |
|
651 |
||
2499
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
652 |
AddSIs [InterI]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
653 |
AddEs [InterD, InterE]; |
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
654 |
|
0bc87b063447
Tidying of proofs. New theorems are enterred immediately into the
paulson
parents:
2031
diff
changeset
|
655 |
|
2912 | 656 |
(*** Image of a set under a function ***) |
657 |
||
658 |
(*Frequently b does not have the syntactic form of f(x).*) |
|
10832 | 659 |
Goalw [image_def] "[| b=f(x); x:A |] ==> b : f`A"; |
5316 | 660 |
by (Blast_tac 1); |
2912 | 661 |
qed "image_eqI"; |
3909 | 662 |
Addsimps [image_eqI]; |
2912 | 663 |
|
664 |
bind_thm ("imageI", refl RS image_eqI); |
|
665 |
||
8025 | 666 |
(*This version's more effective when we already have the required x*) |
10832 | 667 |
Goalw [image_def] "[| x:A; b=f(x) |] ==> b : f`A"; |
8025 | 668 |
by (Blast_tac 1); |
669 |
qed "rev_image_eqI"; |
|
670 |
||
2912 | 671 |
(*The eta-expansion gives variable-name preservation.*) |
5316 | 672 |
val major::prems = Goalw [image_def] |
10832 | 673 |
"[| b : (%x. f(x))`A; !!x.[| b=f(x); x:A |] ==> P |] ==> P"; |
2912 | 674 |
by (rtac (major RS CollectD RS bexE) 1); |
675 |
by (REPEAT (ares_tac prems 1)); |
|
676 |
qed "imageE"; |
|
677 |
||
678 |
AddIs [image_eqI]; |
|
679 |
AddSEs [imageE]; |
|
680 |
||
10832 | 681 |
Goal "f`(A Un B) = f`A Un f`B"; |
2935 | 682 |
by (Blast_tac 1); |
2912 | 683 |
qed "image_Un"; |
684 |
||
10832 | 685 |
Goal "(z : f`A) = (EX x:A. z = f x)"; |
3960 | 686 |
by (Blast_tac 1); |
687 |
qed "image_iff"; |
|
688 |
||
4523 | 689 |
(*This rewrite rule would confuse users if made default.*) |
10832 | 690 |
Goal "(f`A <= B) = (ALL x:A. f(x): B)"; |
4523 | 691 |
by (Blast_tac 1); |
692 |
qed "image_subset_iff"; |
|
693 |
||
11007 | 694 |
Goal "B <= f ` A = (? AA. AA <= A & B = f ` AA)"; |
695 |
by Safe_tac; |
|
696 |
by (Fast_tac 2); |
|
697 |
by (res_inst_tac [("x","{a. a : A & f a : B}")] exI 1); |
|
698 |
by (Fast_tac 1); |
|
699 |
qed "subset_image_iff"; |
|
700 |
||
4523 | 701 |
(*Replaces the three steps subsetI, imageE, hyp_subst_tac, but breaks too |
702 |
many existing proofs.*) |
|
10832 | 703 |
val prems = Goal "(!!x. x:A ==> f(x) : B) ==> f`A <= B"; |
4510 | 704 |
by (blast_tac (claset() addIs prems) 1); |
705 |
qed "image_subsetI"; |
|
706 |
||
2912 | 707 |
(*** Range of a function -- just a translation for image! ***) |
708 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
709 |
Goal "b=f(x) ==> b : range(f)"; |
2912 | 710 |
by (EVERY1 [etac image_eqI, rtac UNIV_I]); |
711 |
bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI)); |
|
712 |
||
713 |
bind_thm ("rangeI", UNIV_I RS imageI); |
|
714 |
||
5316 | 715 |
val [major,minor] = Goal |
3842 | 716 |
"[| b : range(%x. f(x)); !!x. b=f(x) ==> P |] ==> P"; |
2912 | 717 |
by (rtac (major RS imageE) 1); |
718 |
by (etac minor 1); |
|
719 |
qed "rangeE"; |
|
10482 | 720 |
AddXEs [rangeE]; |
2912 | 721 |
|
1776
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
722 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
723 |
(*** Set reasoning tools ***) |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
724 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
725 |
|
3912 | 726 |
(** Rewrite rules for boolean case-splitting: faster than |
4830 | 727 |
addsplits[split_if] |
3912 | 728 |
**) |
729 |
||
4830 | 730 |
bind_thm ("split_if_eq1", read_instantiate [("P", "%x. x = ?b")] split_if); |
731 |
bind_thm ("split_if_eq2", read_instantiate [("P", "%x. ?a = x")] split_if); |
|
3912 | 732 |
|
5237 | 733 |
(*Split ifs on either side of the membership relation. |
734 |
Not for Addsimps -- can cause goals to blow up!*) |
|
9969 | 735 |
bind_thm ("split_if_mem1", inst "P" "%x. x : ?b" split_if); |
736 |
bind_thm ("split_if_mem2", inst "P" "%x. ?a : x" split_if); |
|
3912 | 737 |
|
9108 | 738 |
bind_thms ("split_ifs", [if_bool_eq_conj, split_if_eq1, split_if_eq2, |
9969 | 739 |
split_if_mem1, split_if_mem2]); |
3912 | 740 |
|
741 |
||
4089 | 742 |
(*Each of these has ALREADY been added to simpset() above.*) |
9108 | 743 |
bind_thms ("mem_simps", [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff, |
744 |
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:
1762
diff
changeset
|
745 |
|
9041 | 746 |
(*Would like to add these, but the existing code only searches for the |
747 |
outer-level constant, which in this case is just "op :"; we instead need |
|
748 |
to use term-nets to associate patterns with rules. Also, if a rule fails to |
|
749 |
apply, then the formula should be kept. |
|
750 |
[("uminus", Compl_iff RS iffD1), ("op -", [Diff_iff RS iffD1]), |
|
751 |
("op Int", [IntD1,IntD2]), |
|
752 |
("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] |
|
753 |
*) |
|
754 |
val mksimps_pairs = |
|
755 |
[("Ball",[bspec])] @ mksimps_pairs; |
|
1776
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
756 |
|
6291 | 757 |
simpset_ref() := simpset() setmksimps (mksimps mksimps_pairs); |
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2935
diff
changeset
|
758 |
|
5256 | 759 |
Addsimps[subset_UNIV, subset_refl]; |
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2935
diff
changeset
|
760 |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2935
diff
changeset
|
761 |
|
8001 | 762 |
(*** The 'proper subset' relation (<) ***) |
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2935
diff
changeset
|
763 |
|
5069 | 764 |
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:
2935
diff
changeset
|
765 |
by (Blast_tac 1); |
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2935
diff
changeset
|
766 |
qed "psubsetI"; |
8913 | 767 |
AddSIs [psubsetI]; |
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2935
diff
changeset
|
768 |
|
9088
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents:
9075
diff
changeset
|
769 |
Goalw [psubset_def] |
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents:
9075
diff
changeset
|
770 |
"(A < insert x B) = (if x:B then A<B else if x:A then A-{x} < B else A<=B)"; |
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents:
9075
diff
changeset
|
771 |
by (asm_simp_tac (simpset() addsimps [subset_insert_iff]) 1); |
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents:
9075
diff
changeset
|
772 |
by (Blast_tac 1); |
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
paulson
parents:
9075
diff
changeset
|
773 |
qed "psubset_insert_iff"; |
4059 | 774 |
|
775 |
bind_thm ("psubset_eq", psubset_def RS meta_eq_to_obj_eq); |
|
6443 | 776 |
|
777 |
bind_thm ("psubset_imp_subset", psubset_eq RS iffD1 RS conjunct1); |
|
778 |
||
779 |
Goal"[| (A::'a set) < B; B <= C |] ==> A < C"; |
|
780 |
by (auto_tac (claset(), simpset() addsimps [psubset_eq])); |
|
781 |
qed "psubset_subset_trans"; |
|
782 |
||
783 |
Goal"[| (A::'a set) <= B; B < C|] ==> A < C"; |
|
784 |
by (auto_tac (claset(), simpset() addsimps [psubset_eq])); |
|
785 |
qed "subset_psubset_trans"; |
|
7717
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7658
diff
changeset
|
786 |
|
8001 | 787 |
Goalw [psubset_def] "A < B ==> EX b. b : (B - A)"; |
788 |
by (Blast_tac 1); |
|
789 |
qed "psubset_imp_ex_mem"; |
|
790 |
||
7717
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7658
diff
changeset
|
791 |
|
9892 | 792 |
(* rulify setup *) |
793 |
||
794 |
Goal "(!!x. x:A ==> P x) == Trueprop (ALL x:A. P x)"; |
|
795 |
by (simp_tac (simpset () addsimps (Ball_def :: thms "atomize")) 1); |
|
796 |
qed "ball_eq"; |
|
7717
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7658
diff
changeset
|
797 |
|
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7658
diff
changeset
|
798 |
local |
9892 | 799 |
val ss = HOL_basic_ss addsimps |
800 |
(Drule.norm_hhf_eq :: map Thm.symmetric (ball_eq :: thms "atomize")); |
|
7717
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7658
diff
changeset
|
801 |
in |
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7658
diff
changeset
|
802 |
|
9892 | 803 |
structure Rulify = RulifyFun |
804 |
(val make_meta = Simplifier.simplify ss |
|
805 |
val full_make_meta = Simplifier.full_simplify ss); |
|
806 |
||
807 |
structure BasicRulify: BASIC_RULIFY = Rulify; |
|
808 |
open BasicRulify; |
|
7717
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7658
diff
changeset
|
809 |
|
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
berghofe
parents:
7658
diff
changeset
|
810 |
end; |