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