3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory
4 Copyright 1991 University of Cambridge
6 Set theory for higher-order logic. A set is simply a predicate.
11 section "Relating predicates and sets";
13 AddIffs [mem_Collect_eq];
15 goal Set.thy "!!a. P(a) ==> a : {x.P(x)}";
19 val prems = goal Set.thy "!!a. a : {x.P(x)} ==> P(a)";
20 by (Asm_full_simp_tac 1);
23 val [prem] = goal Set.thy "[| !!x. (x:A) = (x:B) |] ==> A = B";
24 by (rtac (prem RS ext RS arg_cong RS box_equals) 1);
25 by (rtac Collect_mem_eq 1);
26 by (rtac Collect_mem_eq 1);
29 val [prem] = goal Set.thy "[| !!x. P(x)=Q(x) |] ==> {x. P(x)} = {x. Q(x)}";
30 by (rtac (prem RS ext RS arg_cong) 1);
33 val CollectE = make_elim CollectD;
39 section "Bounded quantifiers";
41 val prems = goalw Set.thy [Ball_def]
42 "[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)";
43 by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
46 val [major,minor] = goalw Set.thy [Ball_def]
47 "[| ! x:A. P(x); x:A |] ==> P(x)";
48 by (rtac (minor RS (major RS spec RS mp)) 1);
51 val major::prems = goalw Set.thy [Ball_def]
52 "[| ! x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q";
53 by (rtac (major RS spec RS impCE) 1);
54 by (REPEAT (eresolve_tac prems 1));
57 (*Takes assumptions ! x:A.P(x) and a:A; creates assumption P(a)*)
58 fun ball_tac i = etac ballE i THEN contr_tac (i+1);
63 val prems = goalw Set.thy [Bex_def]
64 "[| P(x); x:A |] ==> ? x:A. P(x)";
65 by (REPEAT (ares_tac (prems @ [exI,conjI]) 1));
68 qed_goal "bexCI" Set.thy
69 "[| ! x:A. ~P(x) ==> P(a); a:A |] ==> ? x:A.P(x)"
72 (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]);
74 val major::prems = goalw Set.thy [Bex_def]
75 "[| ? x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q";
76 by (rtac (major RS exE) 1);
77 by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1));
83 (*Trival rewrite rule; (! x:A.P)=P holds only if A is nonempty!*)
84 goalw Set.thy [Ball_def] "(! x:A. True) = True";
88 (*Dual form for existentials*)
89 goalw Set.thy [Bex_def] "(? x:A. False) = False";
93 Addsimps [ball_True, bex_False];
95 (** Congruence rules **)
97 val prems = goal Set.thy
98 "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \
99 \ (! x:A. P(x)) = (! x:B. Q(x))";
100 by (resolve_tac (prems RL [ssubst]) 1);
101 by (REPEAT (ares_tac [ballI,iffI] 1
102 ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1));
105 val prems = goal Set.thy
106 "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \
107 \ (? x:A. P(x)) = (? x:B. Q(x))";
108 by (resolve_tac (prems RL [ssubst]) 1);
109 by (REPEAT (etac bexE 1
110 ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1));
115 val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B";
116 by (REPEAT (ares_tac (prems @ [ballI]) 1));
119 (*Rule in Modus Ponens style*)
120 val major::prems = goalw Set.thy [subset_def] "[| A <= B; c:A |] ==> c:B";
121 by (rtac (major RS bspec) 1);
122 by (resolve_tac prems 1);
125 (*The same, with reversed premises for use with etac -- cf rev_mp*)
126 qed_goal "rev_subsetD" Set.thy "[| c:A; A <= B |] ==> c:B"
127 (fn prems=> [ (REPEAT (resolve_tac (prems@[subsetD]) 1)) ]);
129 (*Converts A<=B to x:A ==> x:B*)
130 fun impOfSubs th = th RSN (2, rev_subsetD);
132 qed_goal "contra_subsetD" Set.thy "!!c. [| A <= B; c ~: B |] ==> c ~: A"
133 (fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
135 qed_goal "rev_contra_subsetD" Set.thy "!!c. [| c ~: B; A <= B |] ==> c ~: A"
136 (fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
138 (*Classical elimination rule*)
139 val major::prems = goalw Set.thy [subset_def]
140 "[| A <= B; c~:A ==> P; c:B ==> P |] ==> P";
141 by (rtac (major RS ballE) 1);
142 by (REPEAT (eresolve_tac prems 1));
145 (*Takes assumptions A<=B; c:A and creates the assumption c:B *)
146 fun set_mp_tac i = etac subsetCE i THEN mp_tac i;
149 AddEs [subsetD, subsetCE];
151 qed_goal "subset_refl" Set.thy "A <= (A::'a set)"
152 (fn _=> [Fast_tac 1]);
154 val prems = goal Set.thy "!!B. [| A<=B; B<=C |] ==> A<=(C::'a set)";
161 (*Anti-symmetry of the subset relation*)
162 val prems = goal Set.thy "[| A <= B; B <= A |] ==> A = (B::'a set)";
163 by (rtac (iffI RS set_ext) 1);
164 by (REPEAT (ares_tac (prems RL [subsetD]) 1));
165 qed "subset_antisym";
166 val equalityI = subset_antisym;
170 (* Equality rules from ZF set theory -- are they appropriate here? *)
171 val prems = goal Set.thy "A = B ==> A<=(B::'a set)";
172 by (resolve_tac (prems RL [subst]) 1);
173 by (rtac subset_refl 1);
176 val prems = goal Set.thy "A = B ==> B<=(A::'a set)";
177 by (resolve_tac (prems RL [subst]) 1);
178 by (rtac subset_refl 1);
181 val prems = goal Set.thy
182 "[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P";
183 by (resolve_tac prems 1);
184 by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1));
187 val major::prems = goal Set.thy
188 "[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P";
189 by (rtac (major RS equalityE) 1);
190 by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1));
193 (*Lemma for creating induction formulae -- for "pattern matching" on p
194 To make the induction hypotheses usable, apply "spec" or "bspec" to
195 put universal quantifiers over the free variables in p. *)
196 val prems = goal Set.thy
197 "[| p:A; !!z. z:A ==> p=z --> R |] ==> R";
199 by (REPEAT (resolve_tac (refl::prems) 1));
200 qed "setup_induction";
203 section "The empty set -- {}";
205 qed_goalw "empty_iff" Set.thy [empty_def] "(c : {}) = False"
206 (fn _ => [ (Fast_tac 1) ]);
208 Addsimps [empty_iff];
210 qed_goal "emptyE" Set.thy "!!a. a:{} ==> P"
211 (fn _ => [Full_simp_tac 1]);
215 qed_goal "empty_subsetI" Set.thy "{} <= A"
216 (fn _ => [ (Fast_tac 1) ]);
218 qed_goal "equals0I" Set.thy "[| !!y. y:A ==> False |] ==> A={}"
220 [ (fast_tac (!claset addIs [prem RS FalseE]) 1) ]);
222 qed_goal "equals0D" Set.thy "!!a. [| A={}; a:A |] ==> P"
223 (fn _ => [ (Fast_tac 1) ]);
225 goal Set.thy "Ball {} P = True";
226 by (simp_tac (HOL_ss addsimps [mem_Collect_eq, Ball_def, empty_def]) 1);
229 goal Set.thy "Bex {} P = False";
230 by (simp_tac (HOL_ss addsimps [mem_Collect_eq, Bex_def, empty_def]) 1);
232 Addsimps [ball_empty, bex_empty];
234 goalw Set.thy [Ball_def] "(!x:A.False) = (A = {})";
237 Addsimps [ball_False];
239 (* The dual is probably not helpful:
240 goal Set.thy "(? x:A.True) = (A ~= {})";
247 section "The Powerset operator -- Pow";
249 qed_goalw "Pow_iff" Set.thy [Pow_def] "(A : Pow(B)) = (A <= B)"
250 (fn _ => [ (Asm_simp_tac 1) ]);
254 qed_goalw "PowI" Set.thy [Pow_def] "!!A B. A <= B ==> A : Pow(B)"
255 (fn _ => [ (etac CollectI 1) ]);
257 qed_goalw "PowD" Set.thy [Pow_def] "!!A B. A : Pow(B) ==> A<=B"
258 (fn _=> [ (etac CollectD 1) ]);
260 val Pow_bottom = empty_subsetI RS PowI; (* {}: Pow(B) *)
261 val Pow_top = subset_refl RS PowI; (* A : Pow(A) *)
264 section "Set complement -- Compl";
266 qed_goalw "Compl_iff" Set.thy [Compl_def] "(c : Compl(A)) = (c~:A)"
267 (fn _ => [ (Fast_tac 1) ]);
269 Addsimps [Compl_iff];
271 val prems = goalw Set.thy [Compl_def]
272 "[| c:A ==> False |] ==> c : Compl(A)";
273 by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1));
276 (*This form, with negated conclusion, works well with the Classical prover.
277 Negated assumptions behave like formulae on the right side of the notional
279 val major::prems = goalw Set.thy [Compl_def]
280 "c : Compl(A) ==> c~:A";
281 by (rtac (major RS CollectD) 1);
284 val ComplE = make_elim ComplD;
290 section "Binary union -- Un";
292 qed_goalw "Un_iff" Set.thy [Un_def] "(c : A Un B) = (c:A | c:B)"
293 (fn _ => [ Fast_tac 1 ]);
297 goal Set.thy "!!c. c:A ==> c : A Un B";
301 goal Set.thy "!!c. c:B ==> c : A Un B";
305 (*Classical introduction rule: no commitment to A vs B*)
306 qed_goal "UnCI" Set.thy "(c~:B ==> c:A) ==> c : A Un B"
309 (REPEAT (ares_tac (prems@[disjCI]) 1)) ]);
311 val major::prems = goalw Set.thy [Un_def]
312 "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P";
313 by (rtac (major RS CollectD RS disjE) 1);
314 by (REPEAT (eresolve_tac prems 1));
321 section "Binary intersection -- Int";
323 qed_goalw "Int_iff" Set.thy [Int_def] "(c : A Int B) = (c:A & c:B)"
324 (fn _ => [ (Fast_tac 1) ]);
328 goal Set.thy "!!c. [| c:A; c:B |] ==> c : A Int B";
332 goal Set.thy "!!c. c : A Int B ==> c:A";
333 by (Asm_full_simp_tac 1);
336 goal Set.thy "!!c. c : A Int B ==> c:B";
337 by (Asm_full_simp_tac 1);
340 val [major,minor] = goal Set.thy
341 "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P";
343 by (rtac (major RS IntD1) 1);
344 by (rtac (major RS IntD2) 1);
350 section "Set difference";
352 qed_goalw "Diff_iff" Set.thy [set_diff_def] "(c : A-B) = (c:A & c~:B)"
353 (fn _ => [ (Fast_tac 1) ]);
357 qed_goal "DiffI" Set.thy "!!c. [| c : A; c ~: B |] ==> c : A - B"
358 (fn _=> [ Asm_simp_tac 1 ]);
360 qed_goal "DiffD1" Set.thy "!!c. c : A - B ==> c : A"
361 (fn _=> [ (Asm_full_simp_tac 1) ]);
363 qed_goal "DiffD2" Set.thy "!!c. [| c : A - B; c : B |] ==> P"
364 (fn _=> [ (Asm_full_simp_tac 1) ]);
366 qed_goal "DiffE" Set.thy "[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P"
368 [ (resolve_tac prems 1),
369 (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ]);
375 section "Augmenting a set -- insert";
377 qed_goalw "insert_iff" Set.thy [insert_def] "a : insert b A = (a=b | a:A)"
378 (fn _ => [Fast_tac 1]);
380 Addsimps [insert_iff];
382 qed_goal "insertI1" Set.thy "a : insert a B"
383 (fn _ => [Simp_tac 1]);
385 qed_goal "insertI2" Set.thy "!!a. a : B ==> a : insert b B"
386 (fn _=> [Asm_simp_tac 1]);
388 qed_goalw "insertE" Set.thy [insert_def]
389 "[| a : insert b A; a=b ==> P; a:A ==> P |] ==> P"
391 [ (rtac (major RS UnE) 1),
392 (REPEAT (eresolve_tac (prems @ [CollectE]) 1)) ]);
394 (*Classical introduction rule*)
395 qed_goal "insertCI" Set.thy "(a~:B ==> a=b) ==> a: insert b B"
398 (REPEAT (ares_tac (prems@[disjCI]) 1)) ]);
403 section "Singletons, using insert";
405 qed_goal "singletonI" Set.thy "a : {a}"
406 (fn _=> [ (rtac insertI1 1) ]);
408 goal Set.thy "!!a. b : {a} ==> b=a";
412 bind_thm ("singletonE", make_elim singletonD);
414 qed_goal "singleton_iff" thy "(b : {a}) = (b=a)"
415 (fn _ => [Fast_tac 1]);
417 goal Set.thy "!!a b. {a}={b} ==> a=b";
418 by (fast_tac (!claset addEs [equalityE]) 1);
419 qed "singleton_inject";
421 (*Redundant? But unlike insertCI, it proves the subgoal immediately!*)
424 AddSDs [singleton_inject];
427 section "The universal set -- UNIV";
429 qed_goal "UNIV_I" Set.thy "x : UNIV"
430 (fn _ => [rtac ComplI 1, etac emptyE 1]);
432 qed_goal "subset_UNIV" Set.thy "A <= UNIV"
433 (fn _ => [rtac subsetI 1, rtac UNIV_I 1]);
436 section "Unions of families -- UNION x:A. B(x) is Union(B``A)";
438 goalw Set.thy [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))";
444 (*The order of the premises presupposes that A is rigid; b may be flexible*)
445 goal Set.thy "!!b. [| a:A; b: B(a) |] ==> b: (UN x:A. B(x))";
449 val major::prems = goalw Set.thy [UNION_def]
450 "[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R";
451 by (rtac (major RS CollectD RS bexE) 1);
452 by (REPEAT (ares_tac prems 1));
458 val prems = goal Set.thy
459 "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \
460 \ (UN x:A. C(x)) = (UN x:B. D(x))";
461 by (REPEAT (etac UN_E 1
462 ORELSE ares_tac ([UN_I,equalityI,subsetI] @
463 (prems RL [equalityD1,equalityD2] RL [subsetD])) 1));
467 section "Intersections of families -- INTER x:A. B(x) is Inter(B``A)";
469 goalw Set.thy [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))";
475 val prems = goalw Set.thy [INTER_def]
476 "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))";
477 by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1));
480 goal Set.thy "!!b. [| b : (INT x:A. B(x)); a:A |] ==> b: B(a)";
484 (*"Classical" elimination -- by the Excluded Middle on a:A *)
485 val major::prems = goalw Set.thy [INTER_def]
486 "[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R";
487 by (rtac (major RS CollectD RS ballE) 1);
488 by (REPEAT (eresolve_tac prems 1));
492 AddEs [INT_D, INT_E];
494 val prems = goal Set.thy
495 "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \
496 \ (INT x:A. C(x)) = (INT x:B. D(x))";
497 by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI]));
498 by (REPEAT (dtac INT_D 1
499 ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1));
503 section "Unions over a type; UNION1(B) = Union(range(B))";
505 goalw Set.thy [UNION1_def] "(b: (UN x. B(x))) = (EX x. b: B(x))";
512 (*The order of the premises presupposes that A is rigid; b may be flexible*)
513 goal Set.thy "!!b. b: B(x) ==> b: (UN x. B(x))";
517 val major::prems = goalw Set.thy [UNION1_def]
518 "[| b : (UN x. B(x)); !!x. b: B(x) ==> R |] ==> R";
519 by (rtac (major RS UN_E) 1);
520 by (REPEAT (ares_tac prems 1));
527 section "Intersections over a type; INTER1(B) = Inter(range(B))";
529 goalw Set.thy [INTER1_def] "(b: (INT x. B(x))) = (ALL x. b: B(x))";
536 val prems = goalw Set.thy [INTER1_def]
537 "(!!x. b: B(x)) ==> b : (INT x. B(x))";
538 by (REPEAT (ares_tac (INT_I::prems) 1));
541 goal Set.thy "!!b. b : (INT x. B(x)) ==> b: B(a)";
542 by (Asm_full_simp_tac 1);
551 goalw Set.thy [Union_def] "(A : Union(C)) = (EX X:C. A:X)";
555 Addsimps [Union_iff];
557 (*The order of the premises presupposes that C is rigid; A may be flexible*)
558 goal Set.thy "!!X. [| X:C; A:X |] ==> A : Union(C)";
562 val major::prems = goalw Set.thy [Union_def]
563 "[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R";
564 by (rtac (major RS UN_E) 1);
565 by (REPEAT (ares_tac prems 1));
574 goalw Set.thy [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)";
578 Addsimps [Inter_iff];
580 val prems = goalw Set.thy [Inter_def]
581 "[| !!X. X:C ==> A:X |] ==> A : Inter(C)";
582 by (REPEAT (ares_tac ([INT_I] @ prems) 1));
585 (*A "destruct" rule -- every X in C contains A as an element, but
586 A:X can hold when X:C does not! This rule is analogous to "spec". *)
587 goal Set.thy "!!X. [| A : Inter(C); X:C |] ==> A:X";
591 (*"Classical" elimination rule -- does not require proving X:C *)
592 val major::prems = goalw Set.thy [Inter_def]
593 "[| A : Inter(C); X~:C ==> R; A:X ==> R |] ==> R";
594 by (rtac (major RS INT_E) 1);
595 by (REPEAT (eresolve_tac prems 1));
599 AddEs [InterD, InterE];
603 (*** Set reasoning tools ***)
606 (*Each of these has ALREADY been added to !simpset above.*)
607 val mem_simps = [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff,
609 UN_iff, UN1_iff, Union_iff,
610 INT_iff, INT1_iff, Inter_iff];
612 (*Not for Addsimps -- it can cause goals to blow up!*)
613 goal Set.thy "(a : (if Q then x else y)) = ((Q --> a:x) & (~Q --> a : y))";
614 by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
617 val mksimps_pairs = ("Ball",[bspec]) :: mksimps_pairs;
619 simpset := !simpset addcongs [ball_cong,bex_cong]
620 setmksimps (mksimps mksimps_pairs);