author | paulson |
Tue, 16 Jul 1996 15:49:46 +0200 | |
changeset 1868 | 836950047d85 |
parent 1841 | 8e5e2fef6d26 |
child 1882 | 67f49e8c4355 |
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 |
||
6 |
For set.thy. Set theory for higher-order logic. A set is simply a predicate. |
|
7 |
*) |
|
8 |
||
9 |
open Set; |
|
10 |
||
1548 | 11 |
section "Relating predicates and sets"; |
12 |
||
13 |
val [prem] = goal Set.thy "P(a) ==> a : {x.P(x)}"; |
|
923 | 14 |
by (rtac (mem_Collect_eq RS ssubst) 1); |
15 |
by (rtac prem 1); |
|
16 |
qed "CollectI"; |
|
17 |
||
18 |
val prems = goal Set.thy "[| a : {x.P(x)} |] ==> P(a)"; |
|
19 |
by (resolve_tac (prems RL [mem_Collect_eq RS subst]) 1); |
|
20 |
qed "CollectD"; |
|
21 |
||
22 |
val [prem] = goal Set.thy "[| !!x. (x:A) = (x:B) |] ==> A = B"; |
|
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 |
||
28 |
val [prem] = goal Set.thy "[| !!x. P(x)=Q(x) |] ==> {x. P(x)} = {x. Q(x)}"; |
|
29 |
by (rtac (prem RS ext RS arg_cong) 1); |
|
30 |
qed "Collect_cong"; |
|
31 |
||
32 |
val CollectE = make_elim CollectD; |
|
33 |
||
1548 | 34 |
section "Bounded quantifiers"; |
923 | 35 |
|
36 |
val prems = goalw Set.thy [Ball_def] |
|
37 |
"[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)"; |
|
38 |
by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); |
|
39 |
qed "ballI"; |
|
40 |
||
41 |
val [major,minor] = goalw Set.thy [Ball_def] |
|
42 |
"[| ! x:A. P(x); x:A |] ==> P(x)"; |
|
43 |
by (rtac (minor RS (major RS spec RS mp)) 1); |
|
44 |
qed "bspec"; |
|
45 |
||
46 |
val major::prems = goalw Set.thy [Ball_def] |
|
47 |
"[| ! x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q"; |
|
48 |
by (rtac (major RS spec RS impCE) 1); |
|
49 |
by (REPEAT (eresolve_tac prems 1)); |
|
50 |
qed "ballE"; |
|
51 |
||
52 |
(*Takes assumptions ! x:A.P(x) and a:A; creates assumption P(a)*) |
|
53 |
fun ball_tac i = etac ballE i THEN contr_tac (i+1); |
|
54 |
||
55 |
val prems = goalw Set.thy [Bex_def] |
|
56 |
"[| P(x); x:A |] ==> ? x:A. P(x)"; |
|
57 |
by (REPEAT (ares_tac (prems @ [exI,conjI]) 1)); |
|
58 |
qed "bexI"; |
|
59 |
||
60 |
qed_goal "bexCI" Set.thy |
|
61 |
"[| ! x:A. ~P(x) ==> P(a); a:A |] ==> ? x:A.P(x)" |
|
62 |
(fn prems=> |
|
63 |
[ (rtac classical 1), |
|
64 |
(REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]); |
|
65 |
||
66 |
val major::prems = goalw Set.thy [Bex_def] |
|
67 |
"[| ? x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q"; |
|
68 |
by (rtac (major RS exE) 1); |
|
69 |
by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)); |
|
70 |
qed "bexE"; |
|
71 |
||
72 |
(*Trival rewrite rule; (! x:A.P)=P holds only if A is nonempty!*) |
|
1618 | 73 |
goal Set.thy "(! x:A. True) = True"; |
923 | 74 |
by (REPEAT (ares_tac [TrueI,ballI,iffI] 1)); |
1816 | 75 |
qed "ball_True"; |
76 |
||
77 |
Addsimps [ball_True]; |
|
923 | 78 |
|
79 |
(** Congruence rules **) |
|
80 |
||
81 |
val prems = goal Set.thy |
|
82 |
"[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ |
|
83 |
\ (! x:A. P(x)) = (! x:B. Q(x))"; |
|
84 |
by (resolve_tac (prems RL [ssubst]) 1); |
|
85 |
by (REPEAT (ares_tac [ballI,iffI] 1 |
|
86 |
ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1)); |
|
87 |
qed "ball_cong"; |
|
88 |
||
89 |
val prems = goal Set.thy |
|
90 |
"[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ |
|
91 |
\ (? x:A. P(x)) = (? x:B. Q(x))"; |
|
92 |
by (resolve_tac (prems RL [ssubst]) 1); |
|
93 |
by (REPEAT (etac bexE 1 |
|
94 |
ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1)); |
|
95 |
qed "bex_cong"; |
|
96 |
||
1548 | 97 |
section "Subsets"; |
923 | 98 |
|
99 |
val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B"; |
|
100 |
by (REPEAT (ares_tac (prems @ [ballI]) 1)); |
|
101 |
qed "subsetI"; |
|
102 |
||
103 |
(*Rule in Modus Ponens style*) |
|
104 |
val major::prems = goalw Set.thy [subset_def] "[| A <= B; c:A |] ==> c:B"; |
|
105 |
by (rtac (major RS bspec) 1); |
|
106 |
by (resolve_tac prems 1); |
|
107 |
qed "subsetD"; |
|
108 |
||
109 |
(*The same, with reversed premises for use with etac -- cf rev_mp*) |
|
110 |
qed_goal "rev_subsetD" Set.thy "[| c:A; A <= B |] ==> c:B" |
|
111 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[subsetD]) 1)) ]); |
|
112 |
||
1841 | 113 |
qed_goal "contra_subsetD" Set.thy "!!c. [| A <= B; c ~: B |] ==> c ~: A" |
114 |
(fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]); |
|
115 |
||
116 |
qed_goal "rev_contra_subsetD" Set.thy "!!c. [| c ~: B; A <= B |] ==> c ~: A" |
|
117 |
(fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]); |
|
118 |
||
923 | 119 |
(*Classical elimination rule*) |
120 |
val major::prems = goalw Set.thy [subset_def] |
|
121 |
"[| A <= B; c~:A ==> P; c:B ==> P |] ==> P"; |
|
122 |
by (rtac (major RS ballE) 1); |
|
123 |
by (REPEAT (eresolve_tac prems 1)); |
|
124 |
qed "subsetCE"; |
|
125 |
||
126 |
(*Takes assumptions A<=B; c:A and creates the assumption c:B *) |
|
127 |
fun set_mp_tac i = etac subsetCE i THEN mp_tac i; |
|
128 |
||
129 |
qed_goal "subset_refl" Set.thy "A <= (A::'a set)" |
|
130 |
(fn _=> [ (REPEAT (ares_tac [subsetI] 1)) ]); |
|
131 |
||
132 |
val prems = goal Set.thy "[| A<=B; B<=C |] ==> A<=(C::'a set)"; |
|
133 |
by (cut_facts_tac prems 1); |
|
134 |
by (REPEAT (ares_tac [subsetI] 1 ORELSE set_mp_tac 1)); |
|
135 |
qed "subset_trans"; |
|
136 |
||
137 |
||
1548 | 138 |
section "Equality"; |
923 | 139 |
|
140 |
(*Anti-symmetry of the subset relation*) |
|
141 |
val prems = goal Set.thy "[| A <= B; B <= A |] ==> A = (B::'a set)"; |
|
142 |
by (rtac (iffI RS set_ext) 1); |
|
143 |
by (REPEAT (ares_tac (prems RL [subsetD]) 1)); |
|
144 |
qed "subset_antisym"; |
|
145 |
val equalityI = subset_antisym; |
|
146 |
||
1762 | 147 |
AddSIs [equalityI]; |
148 |
||
923 | 149 |
(* Equality rules from ZF set theory -- are they appropriate here? *) |
150 |
val prems = goal Set.thy "A = B ==> A<=(B::'a set)"; |
|
151 |
by (resolve_tac (prems RL [subst]) 1); |
|
152 |
by (rtac subset_refl 1); |
|
153 |
qed "equalityD1"; |
|
154 |
||
155 |
val prems = goal Set.thy "A = B ==> B<=(A::'a set)"; |
|
156 |
by (resolve_tac (prems RL [subst]) 1); |
|
157 |
by (rtac subset_refl 1); |
|
158 |
qed "equalityD2"; |
|
159 |
||
160 |
val prems = goal Set.thy |
|
161 |
"[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P"; |
|
162 |
by (resolve_tac prems 1); |
|
163 |
by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1)); |
|
164 |
qed "equalityE"; |
|
165 |
||
166 |
val major::prems = goal Set.thy |
|
167 |
"[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P"; |
|
168 |
by (rtac (major RS equalityE) 1); |
|
169 |
by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)); |
|
170 |
qed "equalityCE"; |
|
171 |
||
172 |
(*Lemma for creating induction formulae -- for "pattern matching" on p |
|
173 |
To make the induction hypotheses usable, apply "spec" or "bspec" to |
|
174 |
put universal quantifiers over the free variables in p. *) |
|
175 |
val prems = goal Set.thy |
|
176 |
"[| p:A; !!z. z:A ==> p=z --> R |] ==> R"; |
|
177 |
by (rtac mp 1); |
|
178 |
by (REPEAT (resolve_tac (refl::prems) 1)); |
|
179 |
qed "setup_induction"; |
|
180 |
||
181 |
||
1548 | 182 |
section "Set complement -- Compl"; |
923 | 183 |
|
184 |
val prems = goalw Set.thy [Compl_def] |
|
185 |
"[| c:A ==> False |] ==> c : Compl(A)"; |
|
186 |
by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); |
|
187 |
qed "ComplI"; |
|
188 |
||
189 |
(*This form, with negated conclusion, works well with the Classical prover. |
|
190 |
Negated assumptions behave like formulae on the right side of the notional |
|
191 |
turnstile...*) |
|
192 |
val major::prems = goalw Set.thy [Compl_def] |
|
193 |
"[| c : Compl(A) |] ==> c~:A"; |
|
194 |
by (rtac (major RS CollectD) 1); |
|
195 |
qed "ComplD"; |
|
196 |
||
197 |
val ComplE = make_elim ComplD; |
|
198 |
||
1640 | 199 |
qed_goal "Compl_iff" Set.thy "(c : Compl(A)) = (c~:A)" |
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1640
diff
changeset
|
200 |
(fn _ => [ (fast_tac (!claset addSIs [ComplI] addSEs [ComplE]) 1) ]); |
1640 | 201 |
|
923 | 202 |
|
1548 | 203 |
section "Binary union -- Un"; |
923 | 204 |
|
205 |
val prems = goalw Set.thy [Un_def] "c:A ==> c : A Un B"; |
|
206 |
by (REPEAT (resolve_tac (prems @ [CollectI,disjI1]) 1)); |
|
207 |
qed "UnI1"; |
|
208 |
||
209 |
val prems = goalw Set.thy [Un_def] "c:B ==> c : A Un B"; |
|
210 |
by (REPEAT (resolve_tac (prems @ [CollectI,disjI2]) 1)); |
|
211 |
qed "UnI2"; |
|
212 |
||
213 |
(*Classical introduction rule: no commitment to A vs B*) |
|
214 |
qed_goal "UnCI" Set.thy "(c~:B ==> c:A) ==> c : A Un B" |
|
215 |
(fn prems=> |
|
216 |
[ (rtac classical 1), |
|
217 |
(REPEAT (ares_tac (prems@[UnI1,notI]) 1)), |
|
218 |
(REPEAT (ares_tac (prems@[UnI2,notE]) 1)) ]); |
|
219 |
||
220 |
val major::prems = goalw Set.thy [Un_def] |
|
221 |
"[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"; |
|
222 |
by (rtac (major RS CollectD RS disjE) 1); |
|
223 |
by (REPEAT (eresolve_tac prems 1)); |
|
224 |
qed "UnE"; |
|
225 |
||
1640 | 226 |
qed_goal "Un_iff" Set.thy "(c : A Un B) = (c:A | c:B)" |
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1640
diff
changeset
|
227 |
(fn _ => [ (fast_tac (!claset addSIs [UnCI] addSEs [UnE]) 1) ]); |
1640 | 228 |
|
923 | 229 |
|
1548 | 230 |
section "Binary intersection -- Int"; |
923 | 231 |
|
232 |
val prems = goalw Set.thy [Int_def] |
|
233 |
"[| c:A; c:B |] ==> c : A Int B"; |
|
234 |
by (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)); |
|
235 |
qed "IntI"; |
|
236 |
||
237 |
val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:A"; |
|
238 |
by (rtac (major RS CollectD RS conjunct1) 1); |
|
239 |
qed "IntD1"; |
|
240 |
||
241 |
val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:B"; |
|
242 |
by (rtac (major RS CollectD RS conjunct2) 1); |
|
243 |
qed "IntD2"; |
|
244 |
||
245 |
val [major,minor] = goal Set.thy |
|
246 |
"[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"; |
|
247 |
by (rtac minor 1); |
|
248 |
by (rtac (major RS IntD1) 1); |
|
249 |
by (rtac (major RS IntD2) 1); |
|
250 |
qed "IntE"; |
|
251 |
||
1640 | 252 |
qed_goal "Int_iff" Set.thy "(c : A Int B) = (c:A & c:B)" |
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1640
diff
changeset
|
253 |
(fn _ => [ (fast_tac (!claset addSIs [IntI] addSEs [IntE]) 1) ]); |
1640 | 254 |
|
923 | 255 |
|
1548 | 256 |
section "Set difference"; |
923 | 257 |
|
258 |
qed_goalw "DiffI" Set.thy [set_diff_def] |
|
259 |
"[| c : A; c ~: B |] ==> c : A - B" |
|
260 |
(fn prems=> [ (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1)) ]); |
|
261 |
||
262 |
qed_goalw "DiffD1" Set.thy [set_diff_def] |
|
263 |
"c : A - B ==> c : A" |
|
264 |
(fn [major]=> [ (rtac (major RS CollectD RS conjunct1) 1) ]); |
|
265 |
||
266 |
qed_goalw "DiffD2" Set.thy [set_diff_def] |
|
267 |
"[| c : A - B; c : B |] ==> P" |
|
268 |
(fn [major,minor]=> |
|
269 |
[rtac (minor RS (major RS CollectD RS conjunct2 RS notE)) 1]); |
|
270 |
||
271 |
qed_goal "DiffE" Set.thy |
|
272 |
"[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P" |
|
273 |
(fn prems=> |
|
274 |
[ (resolve_tac prems 1), |
|
275 |
(REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ]); |
|
276 |
||
277 |
qed_goal "Diff_iff" Set.thy "(c : A-B) = (c:A & c~:B)" |
|
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1640
diff
changeset
|
278 |
(fn _ => [ (fast_tac (!claset addSIs [DiffI] addSEs [DiffE]) 1) ]); |
923 | 279 |
|
1548 | 280 |
section "The empty set -- {}"; |
923 | 281 |
|
282 |
qed_goalw "emptyE" Set.thy [empty_def] "a:{} ==> P" |
|
283 |
(fn [prem] => [rtac (prem RS CollectD RS FalseE) 1]); |
|
284 |
||
285 |
qed_goal "empty_subsetI" Set.thy "{} <= A" |
|
286 |
(fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]); |
|
287 |
||
288 |
qed_goal "equals0I" Set.thy "[| !!y. y:A ==> False |] ==> A={}" |
|
289 |
(fn prems=> |
|
290 |
[ (REPEAT (ares_tac (prems@[empty_subsetI,subsetI,equalityI]) 1 |
|
291 |
ORELSE eresolve_tac (prems RL [FalseE]) 1)) ]); |
|
292 |
||
293 |
qed_goal "equals0D" Set.thy "[| A={}; a:A |] ==> P" |
|
294 |
(fn [major,minor]=> |
|
295 |
[ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]); |
|
296 |
||
1640 | 297 |
qed_goal "empty_iff" Set.thy "(c : {}) = False" |
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1640
diff
changeset
|
298 |
(fn _ => [ (fast_tac (!claset addSEs [emptyE]) 1) ]); |
1640 | 299 |
|
1816 | 300 |
goal Set.thy "Ball {} P = True"; |
301 |
by (simp_tac (HOL_ss addsimps [mem_Collect_eq, Ball_def, empty_def]) 1); |
|
302 |
qed "ball_empty"; |
|
303 |
||
304 |
goal Set.thy "Bex {} P = False"; |
|
305 |
by (simp_tac (HOL_ss addsimps [mem_Collect_eq, Bex_def, empty_def]) 1); |
|
306 |
qed "bex_empty"; |
|
307 |
Addsimps [ball_empty, bex_empty]; |
|
308 |
||
923 | 309 |
|
1548 | 310 |
section "Augmenting a set -- insert"; |
923 | 311 |
|
312 |
qed_goalw "insertI1" Set.thy [insert_def] "a : insert a B" |
|
313 |
(fn _ => [rtac (CollectI RS UnI1) 1, rtac refl 1]); |
|
314 |
||
315 |
qed_goalw "insertI2" Set.thy [insert_def] "a : B ==> a : insert b B" |
|
316 |
(fn [prem]=> [ (rtac (prem RS UnI2) 1) ]); |
|
317 |
||
318 |
qed_goalw "insertE" Set.thy [insert_def] |
|
319 |
"[| a : insert b A; a=b ==> P; a:A ==> P |] ==> P" |
|
320 |
(fn major::prems=> |
|
321 |
[ (rtac (major RS UnE) 1), |
|
322 |
(REPEAT (eresolve_tac (prems @ [CollectE]) 1)) ]); |
|
323 |
||
324 |
qed_goal "insert_iff" Set.thy "a : insert b A = (a=b | a:A)" |
|
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1640
diff
changeset
|
325 |
(fn _ => [fast_tac (!claset addIs [insertI1,insertI2] addSEs [insertE]) 1]); |
923 | 326 |
|
327 |
(*Classical introduction rule*) |
|
328 |
qed_goal "insertCI" Set.thy "(a~:B ==> a=b) ==> a: insert b B" |
|
329 |
(fn [prem]=> |
|
330 |
[ (rtac (disjCI RS (insert_iff RS iffD2)) 1), |
|
331 |
(etac prem 1) ]); |
|
332 |
||
1548 | 333 |
section "Singletons, using insert"; |
923 | 334 |
|
335 |
qed_goal "singletonI" Set.thy "a : {a}" |
|
336 |
(fn _=> [ (rtac insertI1 1) ]); |
|
337 |
||
338 |
goalw Set.thy [insert_def] "!!a. b : {a} ==> b=a"; |
|
1760
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1640
diff
changeset
|
339 |
by (fast_tac (!claset addSEs [emptyE,CollectE,UnE]) 1); |
923 | 340 |
qed "singletonD"; |
341 |
||
1776
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
342 |
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
|
343 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
344 |
qed_goal "singleton_iff" thy "(b : {a}) = (b=a)" (fn _ => [ |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
345 |
rtac iffI 1, |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
346 |
etac singletonD 1, |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
347 |
hyp_subst_tac 1, |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
348 |
rtac singletonI 1]); |
923 | 349 |
|
350 |
val [major] = goal Set.thy "{a}={b} ==> a=b"; |
|
351 |
by (rtac (major RS equalityD1 RS subsetD RS singletonD) 1); |
|
352 |
by (rtac singletonI 1); |
|
353 |
qed "singleton_inject"; |
|
354 |
||
1531 | 355 |
|
1548 | 356 |
section "The universal set -- UNIV"; |
1531 | 357 |
|
358 |
qed_goal "subset_UNIV" Set.thy "A <= UNIV" |
|
359 |
(fn _ => [rtac subsetI 1, rtac ComplI 1, etac emptyE 1]); |
|
360 |
||
361 |
||
1548 | 362 |
section "Unions of families -- UNION x:A. B(x) is Union(B``A)"; |
923 | 363 |
|
364 |
(*The order of the premises presupposes that A is rigid; b may be flexible*) |
|
365 |
val prems = goalw Set.thy [UNION_def] |
|
366 |
"[| a:A; b: B(a) |] ==> b: (UN x:A. B(x))"; |
|
367 |
by (REPEAT (resolve_tac (prems @ [bexI,CollectI]) 1)); |
|
368 |
qed "UN_I"; |
|
369 |
||
370 |
val major::prems = goalw Set.thy [UNION_def] |
|
371 |
"[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R"; |
|
372 |
by (rtac (major RS CollectD RS bexE) 1); |
|
373 |
by (REPEAT (ares_tac prems 1)); |
|
374 |
qed "UN_E"; |
|
375 |
||
376 |
val prems = goal Set.thy |
|
377 |
"[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ |
|
378 |
\ (UN x:A. C(x)) = (UN x:B. D(x))"; |
|
379 |
by (REPEAT (etac UN_E 1 |
|
380 |
ORELSE ares_tac ([UN_I,equalityI,subsetI] @ |
|
1465 | 381 |
(prems RL [equalityD1,equalityD2] RL [subsetD])) 1)); |
923 | 382 |
qed "UN_cong"; |
383 |
||
384 |
||
1548 | 385 |
section "Intersections of families -- INTER x:A. B(x) is Inter(B``A)"; |
923 | 386 |
|
387 |
val prems = goalw Set.thy [INTER_def] |
|
388 |
"(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))"; |
|
389 |
by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1)); |
|
390 |
qed "INT_I"; |
|
391 |
||
392 |
val major::prems = goalw Set.thy [INTER_def] |
|
393 |
"[| b : (INT x:A. B(x)); a:A |] ==> b: B(a)"; |
|
394 |
by (rtac (major RS CollectD RS bspec) 1); |
|
395 |
by (resolve_tac prems 1); |
|
396 |
qed "INT_D"; |
|
397 |
||
398 |
(*"Classical" elimination -- by the Excluded Middle on a:A *) |
|
399 |
val major::prems = goalw Set.thy [INTER_def] |
|
400 |
"[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R"; |
|
401 |
by (rtac (major RS CollectD RS ballE) 1); |
|
402 |
by (REPEAT (eresolve_tac prems 1)); |
|
403 |
qed "INT_E"; |
|
404 |
||
405 |
val prems = goal Set.thy |
|
406 |
"[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ |
|
407 |
\ (INT x:A. C(x)) = (INT x:B. D(x))"; |
|
408 |
by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI])); |
|
409 |
by (REPEAT (dtac INT_D 1 |
|
410 |
ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1)); |
|
411 |
qed "INT_cong"; |
|
412 |
||
413 |
||
1548 | 414 |
section "Unions over a type; UNION1(B) = Union(range(B))"; |
923 | 415 |
|
416 |
(*The order of the premises presupposes that A is rigid; b may be flexible*) |
|
417 |
val prems = goalw Set.thy [UNION1_def] |
|
418 |
"b: B(x) ==> b: (UN x. B(x))"; |
|
419 |
by (REPEAT (resolve_tac (prems @ [TrueI, CollectI RS UN_I]) 1)); |
|
420 |
qed "UN1_I"; |
|
421 |
||
422 |
val major::prems = goalw Set.thy [UNION1_def] |
|
423 |
"[| b : (UN x. B(x)); !!x. b: B(x) ==> R |] ==> R"; |
|
424 |
by (rtac (major RS UN_E) 1); |
|
425 |
by (REPEAT (ares_tac prems 1)); |
|
426 |
qed "UN1_E"; |
|
427 |
||
428 |
||
1548 | 429 |
section "Intersections over a type; INTER1(B) = Inter(range(B))"; |
923 | 430 |
|
431 |
val prems = goalw Set.thy [INTER1_def] |
|
432 |
"(!!x. b: B(x)) ==> b : (INT x. B(x))"; |
|
433 |
by (REPEAT (ares_tac (INT_I::prems) 1)); |
|
434 |
qed "INT1_I"; |
|
435 |
||
436 |
val [major] = goalw Set.thy [INTER1_def] |
|
437 |
"b : (INT x. B(x)) ==> b: B(a)"; |
|
438 |
by (rtac (TrueI RS (CollectI RS (major RS INT_D))) 1); |
|
439 |
qed "INT1_D"; |
|
440 |
||
1548 | 441 |
section "Union"; |
923 | 442 |
|
443 |
(*The order of the premises presupposes that C is rigid; A may be flexible*) |
|
444 |
val prems = goalw Set.thy [Union_def] |
|
445 |
"[| X:C; A:X |] ==> A : Union(C)"; |
|
446 |
by (REPEAT (resolve_tac (prems @ [UN_I]) 1)); |
|
447 |
qed "UnionI"; |
|
448 |
||
449 |
val major::prems = goalw Set.thy [Union_def] |
|
450 |
"[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R"; |
|
451 |
by (rtac (major RS UN_E) 1); |
|
452 |
by (REPEAT (ares_tac prems 1)); |
|
453 |
qed "UnionE"; |
|
454 |
||
1548 | 455 |
section "Inter"; |
923 | 456 |
|
457 |
val prems = goalw Set.thy [Inter_def] |
|
458 |
"[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; |
|
459 |
by (REPEAT (ares_tac ([INT_I] @ prems) 1)); |
|
460 |
qed "InterI"; |
|
461 |
||
462 |
(*A "destruct" rule -- every X in C contains A as an element, but |
|
463 |
A:X can hold when X:C does not! This rule is analogous to "spec". *) |
|
464 |
val major::prems = goalw Set.thy [Inter_def] |
|
465 |
"[| A : Inter(C); X:C |] ==> A:X"; |
|
466 |
by (rtac (major RS INT_D) 1); |
|
467 |
by (resolve_tac prems 1); |
|
468 |
qed "InterD"; |
|
469 |
||
470 |
(*"Classical" elimination rule -- does not require proving X:C *) |
|
471 |
val major::prems = goalw Set.thy [Inter_def] |
|
472 |
"[| A : Inter(C); A:X ==> R; X~:C ==> R |] ==> R"; |
|
473 |
by (rtac (major RS INT_E) 1); |
|
474 |
by (REPEAT (eresolve_tac prems 1)); |
|
475 |
qed "InterE"; |
|
476 |
||
1548 | 477 |
section "The Powerset operator -- Pow"; |
923 | 478 |
|
479 |
qed_goalw "PowI" Set.thy [Pow_def] "!!A B. A <= B ==> A : Pow(B)" |
|
480 |
(fn _ => [ (etac CollectI 1) ]); |
|
481 |
||
482 |
qed_goalw "PowD" Set.thy [Pow_def] "!!A B. A : Pow(B) ==> A<=B" |
|
483 |
(fn _=> [ (etac CollectD 1) ]); |
|
484 |
||
485 |
val Pow_bottom = empty_subsetI RS PowI; (* {}: Pow(B) *) |
|
486 |
val Pow_top = subset_refl RS PowI; (* A : Pow(A) *) |
|
1776
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
487 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
488 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
489 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
490 |
(*** Set reasoning tools ***) |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
491 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
492 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
493 |
val mem_simps = [ Un_iff, Int_iff, Compl_iff, Diff_iff, singleton_iff, |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
494 |
mem_Collect_eq]; |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
495 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
496 |
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:
1762
diff
changeset
|
497 |
|
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
498 |
simpset := !simpset addsimps mem_simps |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
499 |
addcongs [ball_cong,bex_cong] |
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
oheimb
parents:
1762
diff
changeset
|
500 |
setmksimps (mksimps mksimps_pairs); |