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