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 Addsimps [Collect_mem_eq];
14 AddIffs [mem_Collect_eq];
16 goal Set.thy "!!a. P(a) ==> a : {x. P(x)}";
20 val prems = goal Set.thy "!!a. a : {x. P(x)} ==> P(a)";
21 by (Asm_full_simp_tac 1);
24 val [prem] = goal Set.thy "[| !!x. (x:A) = (x:B) |] ==> A = B";
25 by (rtac (prem RS ext RS arg_cong RS box_equals) 1);
26 by (rtac Collect_mem_eq 1);
27 by (rtac Collect_mem_eq 1);
30 val [prem] = goal Set.thy "[| !!x. P(x)=Q(x) |] ==> {x. P(x)} = {x. Q(x)}";
31 by (rtac (prem RS ext RS arg_cong) 1);
34 val CollectE = make_elim CollectD;
40 section "Bounded quantifiers";
42 val prems = goalw Set.thy [Ball_def]
43 "[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)";
44 by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
47 val [major,minor] = goalw Set.thy [Ball_def]
48 "[| ! x:A. P(x); x:A |] ==> P(x)";
49 by (rtac (minor RS (major RS spec RS mp)) 1);
52 val major::prems = goalw Set.thy [Ball_def]
53 "[| ! x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q";
54 by (rtac (major RS spec RS impCE) 1);
55 by (REPEAT (eresolve_tac prems 1));
58 (*Takes assumptions ! x:A.P(x) and a:A; creates assumption P(a)*)
59 fun ball_tac i = etac ballE i THEN contr_tac (i+1);
64 val prems = goalw Set.thy [Bex_def]
65 "[| P(x); x:A |] ==> ? x:A. P(x)";
66 by (REPEAT (ares_tac (prems @ [exI,conjI]) 1));
69 qed_goal "bexCI" Set.thy
70 "[| ! x:A. ~P(x) ==> P(a); a:A |] ==> ? x:A. P(x)"
73 (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ]);
75 val major::prems = goalw Set.thy [Bex_def]
76 "[| ? x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q";
77 by (rtac (major RS exE) 1);
78 by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1));
84 (*Trival rewrite rule*)
85 goal Set.thy "(! x:A. P) = ((? x. x:A) --> P)";
86 by (simp_tac (simpset() addsimps [Ball_def]) 1);
89 (*Dual form for existentials*)
90 goal Set.thy "(? x:A. P) = ((? x. x:A) & P)";
91 by (simp_tac (simpset() addsimps [Bex_def]) 1);
94 Addsimps [ball_triv, bex_triv];
96 (** Congruence rules **)
98 val prems = goal Set.thy
99 "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \
100 \ (! x:A. P(x)) = (! x:B. Q(x))";
101 by (resolve_tac (prems RL [ssubst]) 1);
102 by (REPEAT (ares_tac [ballI,iffI] 1
103 ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1));
106 val prems = goal Set.thy
107 "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \
108 \ (? x:A. P(x)) = (? x:B. Q(x))";
109 by (resolve_tac (prems RL [ssubst]) 1);
110 by (REPEAT (etac bexE 1
111 ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1));
116 val prems = goalw Set.thy [subset_def] "(!!x. x:A ==> x:B) ==> A <= B";
117 by (REPEAT (ares_tac (prems @ [ballI]) 1));
120 Blast.overload ("op <=", domain_type); (*The <= relation is overloaded*)
122 (*While (:) is not, its type must be kept
123 for overloading of = to work.*)
124 Blast.overload ("op :", domain_type);
125 seq (fn a => Blast.overload (a, HOLogic.dest_setT o domain_type))
127 (*need UNION, INTER also?*)
130 (*Rule in Modus Ponens style*)
131 val major::prems = goalw Set.thy [subset_def] "[| A <= B; c:A |] ==> c:B";
132 by (rtac (major RS bspec) 1);
133 by (resolve_tac prems 1);
136 (*The same, with reversed premises for use with etac -- cf rev_mp*)
137 qed_goal "rev_subsetD" Set.thy "[| c:A; A <= B |] ==> c:B"
138 (fn prems=> [ (REPEAT (resolve_tac (prems@[subsetD]) 1)) ]);
140 (*Converts A<=B to x:A ==> x:B*)
141 fun impOfSubs th = th RSN (2, rev_subsetD);
143 qed_goal "contra_subsetD" Set.thy "!!c. [| A <= B; c ~: B |] ==> c ~: A"
144 (fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
146 qed_goal "rev_contra_subsetD" Set.thy "!!c. [| c ~: B; A <= B |] ==> c ~: A"
147 (fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
149 (*Classical elimination rule*)
150 val major::prems = goalw Set.thy [subset_def]
151 "[| A <= B; c~:A ==> P; c:B ==> P |] ==> P";
152 by (rtac (major RS ballE) 1);
153 by (REPEAT (eresolve_tac prems 1));
156 (*Takes assumptions A<=B; c:A and creates the assumption c:B *)
157 fun set_mp_tac i = etac subsetCE i THEN mp_tac i;
160 AddEs [subsetD, subsetCE];
162 qed_goal "subset_refl" Set.thy "A <= (A::'a set)"
163 (fn _=> [Fast_tac 1]); (*Blast_tac would try order_refl and fail*)
165 val prems = goal Set.thy "!!B. [| A<=B; B<=C |] ==> A<=(C::'a set)";
172 (*Anti-symmetry of the subset relation*)
173 val prems = goal Set.thy "[| A <= B; B <= A |] ==> A = (B::'a set)";
174 by (rtac (iffI RS set_ext) 1);
175 by (REPEAT (ares_tac (prems RL [subsetD]) 1));
176 qed "subset_antisym";
177 val equalityI = subset_antisym;
181 (* Equality rules from ZF set theory -- are they appropriate here? *)
182 val prems = goal Set.thy "A = B ==> A<=(B::'a set)";
183 by (resolve_tac (prems RL [subst]) 1);
184 by (rtac subset_refl 1);
187 val prems = goal Set.thy "A = B ==> B<=(A::'a set)";
188 by (resolve_tac (prems RL [subst]) 1);
189 by (rtac subset_refl 1);
192 val prems = goal Set.thy
193 "[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P";
194 by (resolve_tac prems 1);
195 by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1));
198 val major::prems = goal Set.thy
199 "[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P";
200 by (rtac (major RS equalityE) 1);
201 by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1));
204 (*Lemma for creating induction formulae -- for "pattern matching" on p
205 To make the induction hypotheses usable, apply "spec" or "bspec" to
206 put universal quantifiers over the free variables in p. *)
207 val prems = goal Set.thy
208 "[| p:A; !!z. z:A ==> p=z --> R |] ==> R";
210 by (REPEAT (resolve_tac (refl::prems) 1));
211 qed "setup_induction";
214 section "The universal set -- UNIV";
216 qed_goalw "UNIV_I" Set.thy [UNIV_def] "x : UNIV"
217 (fn _ => [rtac CollectI 1, rtac TrueI 1]);
221 qed_goal "subset_UNIV" Set.thy "A <= UNIV"
222 (fn _ => [rtac subsetI 1, rtac UNIV_I 1]);
224 (** Eta-contracting these two rules (to remove P) causes them to be ignored
225 because of their interaction with congruence rules. **)
227 goalw Set.thy [Ball_def] "Ball UNIV P = All P";
231 goalw Set.thy [Bex_def] "Bex UNIV P = Ex P";
234 Addsimps [ball_UNIV, bex_UNIV];
237 section "The empty set -- {}";
239 qed_goalw "empty_iff" Set.thy [empty_def] "(c : {}) = False"
240 (fn _ => [ (Blast_tac 1) ]);
242 Addsimps [empty_iff];
244 qed_goal "emptyE" Set.thy "!!a. a:{} ==> P"
245 (fn _ => [Full_simp_tac 1]);
249 qed_goal "empty_subsetI" Set.thy "{} <= A"
250 (fn _ => [ (Blast_tac 1) ]);
252 qed_goal "equals0I" Set.thy "[| !!y. y:A ==> False |] ==> A={}"
254 [ (blast_tac (claset() addIs [prem RS FalseE]) 1) ]);
256 qed_goal "equals0D" Set.thy "!!a. [| A={}; a:A |] ==> P"
257 (fn _ => [ (Blast_tac 1) ]);
259 goalw Set.thy [Ball_def] "Ball {} P = True";
263 goalw Set.thy [Bex_def] "Bex {} P = False";
266 Addsimps [ball_empty, bex_empty];
268 goal thy "UNIV ~= {}";
269 by (blast_tac (claset() addEs [equalityE]) 1);
270 qed "UNIV_not_empty";
271 AddIffs [UNIV_not_empty];
275 section "The Powerset operator -- Pow";
277 qed_goalw "Pow_iff" Set.thy [Pow_def] "(A : Pow(B)) = (A <= B)"
278 (fn _ => [ (Asm_simp_tac 1) ]);
282 qed_goalw "PowI" Set.thy [Pow_def] "!!A B. A <= B ==> A : Pow(B)"
283 (fn _ => [ (etac CollectI 1) ]);
285 qed_goalw "PowD" Set.thy [Pow_def] "!!A B. A : Pow(B) ==> A<=B"
286 (fn _=> [ (etac CollectD 1) ]);
288 val Pow_bottom = empty_subsetI RS PowI; (* {}: Pow(B) *)
289 val Pow_top = subset_refl RS PowI; (* A : Pow(A) *)
292 section "Set complement -- Compl";
294 qed_goalw "Compl_iff" Set.thy [Compl_def] "(c : Compl(A)) = (c~:A)"
295 (fn _ => [ (Blast_tac 1) ]);
297 Addsimps [Compl_iff];
299 val prems = goalw Set.thy [Compl_def]
300 "[| c:A ==> False |] ==> c : Compl(A)";
301 by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1));
304 (*This form, with negated conclusion, works well with the Classical prover.
305 Negated assumptions behave like formulae on the right side of the notional
307 val major::prems = goalw Set.thy [Compl_def]
308 "c : Compl(A) ==> c~:A";
309 by (rtac (major RS CollectD) 1);
312 val ComplE = make_elim ComplD;
318 section "Binary union -- Un";
320 qed_goalw "Un_iff" Set.thy [Un_def] "(c : A Un B) = (c:A | c:B)"
321 (fn _ => [ Blast_tac 1 ]);
325 goal Set.thy "!!c. c:A ==> c : A Un B";
329 goal Set.thy "!!c. c:B ==> c : A Un B";
333 (*Classical introduction rule: no commitment to A vs B*)
334 qed_goal "UnCI" Set.thy "(c~:B ==> c:A) ==> c : A Un B"
337 (REPEAT (ares_tac (prems@[disjCI]) 1)) ]);
339 val major::prems = goalw Set.thy [Un_def]
340 "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P";
341 by (rtac (major RS CollectD RS disjE) 1);
342 by (REPEAT (eresolve_tac prems 1));
349 section "Binary intersection -- Int";
351 qed_goalw "Int_iff" Set.thy [Int_def] "(c : A Int B) = (c:A & c:B)"
352 (fn _ => [ (Blast_tac 1) ]);
356 goal Set.thy "!!c. [| c:A; c:B |] ==> c : A Int B";
360 goal Set.thy "!!c. c : A Int B ==> c:A";
361 by (Asm_full_simp_tac 1);
364 goal Set.thy "!!c. c : A Int B ==> c:B";
365 by (Asm_full_simp_tac 1);
368 val [major,minor] = goal Set.thy
369 "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P";
371 by (rtac (major RS IntD1) 1);
372 by (rtac (major RS IntD2) 1);
378 section "Set difference";
380 qed_goalw "Diff_iff" Set.thy [set_diff_def] "(c : A-B) = (c:A & c~:B)"
381 (fn _ => [ (Blast_tac 1) ]);
385 qed_goal "DiffI" Set.thy "!!c. [| c : A; c ~: B |] ==> c : A - B"
386 (fn _=> [ Asm_simp_tac 1 ]);
388 qed_goal "DiffD1" Set.thy "!!c. c : A - B ==> c : A"
389 (fn _=> [ (Asm_full_simp_tac 1) ]);
391 qed_goal "DiffD2" Set.thy "!!c. [| c : A - B; c : B |] ==> P"
392 (fn _=> [ (Asm_full_simp_tac 1) ]);
394 qed_goal "DiffE" Set.thy "[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P"
396 [ (resolve_tac prems 1),
397 (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ]);
403 section "Augmenting a set -- insert";
405 qed_goalw "insert_iff" Set.thy [insert_def] "a : insert b A = (a=b | a:A)"
406 (fn _ => [Blast_tac 1]);
408 Addsimps [insert_iff];
410 qed_goal "insertI1" Set.thy "a : insert a B"
411 (fn _ => [Simp_tac 1]);
413 qed_goal "insertI2" Set.thy "!!a. a : B ==> a : insert b B"
414 (fn _=> [Asm_simp_tac 1]);
416 qed_goalw "insertE" Set.thy [insert_def]
417 "[| a : insert b A; a=b ==> P; a:A ==> P |] ==> P"
419 [ (rtac (major RS UnE) 1),
420 (REPEAT (eresolve_tac (prems @ [CollectE]) 1)) ]);
422 (*Classical introduction rule*)
423 qed_goal "insertCI" Set.thy "(a~:B ==> a=b) ==> a: insert b B"
426 (REPEAT (ares_tac (prems@[disjCI]) 1)) ]);
431 section "Singletons, using insert";
433 qed_goal "singletonI" Set.thy "a : {a}"
434 (fn _=> [ (rtac insertI1 1) ]);
436 goal Set.thy "!!a. b : {a} ==> b=a";
440 bind_thm ("singletonE", make_elim singletonD);
442 qed_goal "singleton_iff" thy "(b : {a}) = (b=a)"
443 (fn _ => [Blast_tac 1]);
445 goal Set.thy "!!a b. {a}={b} ==> a=b";
446 by (blast_tac (claset() addEs [equalityE]) 1);
447 qed "singleton_inject";
449 (*Redundant? But unlike insertCI, it proves the subgoal immediately!*)
451 AddSDs [singleton_inject];
454 goal Set.thy "{x. x=a} = {a}";
456 qed "singleton_conv";
457 Addsimps [singleton_conv];
460 section "Unions of families -- UNION x:A. B(x) is Union(B``A)";
462 goalw Set.thy [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))";
468 (*The order of the premises presupposes that A is rigid; b may be flexible*)
469 goal Set.thy "!!b. [| a:A; b: B(a) |] ==> b: (UN x:A. B(x))";
473 val major::prems = goalw Set.thy [UNION_def]
474 "[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R";
475 by (rtac (major RS CollectD RS bexE) 1);
476 by (REPEAT (ares_tac prems 1));
482 val prems = goal Set.thy
483 "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \
484 \ (UN x:A. C(x)) = (UN x:B. D(x))";
485 by (REPEAT (etac UN_E 1
486 ORELSE ares_tac ([UN_I,equalityI,subsetI] @
487 (prems RL [equalityD1,equalityD2] RL [subsetD])) 1));
491 section "Intersections of families -- INTER x:A. B(x) is Inter(B``A)";
493 goalw Set.thy [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))";
499 val prems = goalw Set.thy [INTER_def]
500 "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))";
501 by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1));
504 goal Set.thy "!!b. [| b : (INT x:A. B(x)); a:A |] ==> b: B(a)";
508 (*"Classical" elimination -- by the Excluded Middle on a:A *)
509 val major::prems = goalw Set.thy [INTER_def]
510 "[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R";
511 by (rtac (major RS CollectD RS ballE) 1);
512 by (REPEAT (eresolve_tac prems 1));
516 AddEs [INT_D, INT_E];
518 val prems = goal Set.thy
519 "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \
520 \ (INT x:A. C(x)) = (INT x:B. D(x))";
521 by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI]));
522 by (REPEAT (dtac INT_D 1
523 ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1));
529 goalw Set.thy [Union_def] "(A : Union(C)) = (EX X:C. A:X)";
533 Addsimps [Union_iff];
535 (*The order of the premises presupposes that C is rigid; A may be flexible*)
536 goal Set.thy "!!X. [| X:C; A:X |] ==> A : Union(C)";
540 val major::prems = goalw Set.thy [Union_def]
541 "[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R";
542 by (rtac (major RS UN_E) 1);
543 by (REPEAT (ares_tac prems 1));
552 goalw Set.thy [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)";
556 Addsimps [Inter_iff];
558 val prems = goalw Set.thy [Inter_def]
559 "[| !!X. X:C ==> A:X |] ==> A : Inter(C)";
560 by (REPEAT (ares_tac ([INT_I] @ prems) 1));
563 (*A "destruct" rule -- every X in C contains A as an element, but
564 A:X can hold when X:C does not! This rule is analogous to "spec". *)
565 goal Set.thy "!!X. [| A : Inter(C); X:C |] ==> A:X";
569 (*"Classical" elimination rule -- does not require proving X:C *)
570 val major::prems = goalw Set.thy [Inter_def]
571 "[| A : Inter(C); X~:C ==> R; A:X ==> R |] ==> R";
572 by (rtac (major RS INT_E) 1);
573 by (REPEAT (eresolve_tac prems 1));
577 AddEs [InterD, InterE];
580 (*** Image of a set under a function ***)
582 (*Frequently b does not have the syntactic form of f(x).*)
583 val prems = goalw thy [image_def] "[| b=f(x); x:A |] ==> b : f``A";
584 by (REPEAT (resolve_tac (prems @ [CollectI,bexI,prem]) 1));
586 Addsimps [image_eqI];
588 bind_thm ("imageI", refl RS image_eqI);
590 (*The eta-expansion gives variable-name preservation.*)
591 val major::prems = goalw thy [image_def]
592 "[| b : (%x. f(x))``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P";
593 by (rtac (major RS CollectD RS bexE) 1);
594 by (REPEAT (ares_tac prems 1));
600 goalw thy [o_def] "(f o g)``r = f``(g``r)";
604 goal thy "f``(A Un B) = f``A Un f``B";
608 goal Set.thy "(z : f``A) = (EX x:A. z = f x)";
613 goal Set.thy "(! x : f `` A. P x) = (! a:A. P(f a))";
617 goal Set.thy "(? x : f `` A. P x) = (? a:A. P(f a))";
621 Addsimps [ball_image,bex_image];
623 (*** Range of a function -- just a translation for image! ***)
625 goal thy "!!b. b=f(x) ==> b : range(f)";
626 by (EVERY1 [etac image_eqI, rtac UNIV_I]);
627 bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI));
629 bind_thm ("rangeI", UNIV_I RS imageI);
631 val [major,minor] = goal thy
632 "[| b : range(%x. f(x)); !!x. b=f(x) ==> P |] ==> P";
633 by (rtac (major RS imageE) 1);
638 (*** Set reasoning tools ***)
641 (** Rewrite rules for boolean case-splitting: faster than
645 bind_thm ("expand_if_eq1", read_instantiate [("P", "%x. x = ?b")] expand_if);
646 bind_thm ("expand_if_eq2", read_instantiate [("P", "%x. ?a = x")] expand_if);
648 bind_thm ("expand_if_mem1",
649 read_instantiate_sg (sign_of Set.thy) [("P", "%x. x : ?b")] expand_if);
650 bind_thm ("expand_if_mem2",
651 read_instantiate_sg (sign_of Set.thy) [("P", "%x. ?a : x")] expand_if);
653 val expand_ifs = [if_bool_eq, expand_if_eq1, expand_if_eq2,
654 expand_if_mem1, expand_if_mem2];
657 (*Each of these has ALREADY been added to simpset() above.*)
658 val mem_simps = [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff,
659 mem_Collect_eq, UN_iff, Union_iff, INT_iff, Inter_iff];
661 (*Not for Addsimps -- it can cause goals to blow up!*)
662 goal Set.thy "(a : (if Q then x else y)) = ((Q --> a:x) & (~Q --> a : y))";
663 by (simp_tac (simpset() addsplits [expand_if]) 1);
666 val mksimps_pairs = ("Ball",[bspec]) :: mksimps_pairs;
668 simpset_ref() := simpset() addcongs [ball_cong,bex_cong]
669 setmksimps (mksimps mksimps_pairs);
671 Addsimps[subset_UNIV, empty_subsetI, subset_refl];
676 goalw Set.thy [psubset_def] "!!A::'a set. [| A <= B; A ~= B |] ==> A<B";
680 goalw Set.thy [psubset_def]
681 "!!x. A < insert x B ==> (x ~: A) & A<=B | x:A & A-{x}<B";
683 qed "psubset_insertD";
685 bind_thm ("psubset_eq", psubset_def RS meta_eq_to_obj_eq);