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