9869
|
1 |
(* Title: HOL/meson_lemmas.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1992 University of Cambridge
|
|
5 |
|
|
6 |
Lemmas for Meson.
|
|
7 |
*)
|
|
8 |
|
|
9 |
(* Generation of contrapositives *)
|
|
10 |
|
|
11 |
(*Inserts negated disjunct after removing the negation; P is a literal*)
|
|
12 |
val [major,minor] = Goal "~P|Q ==> ((~P==>P) ==> Q)";
|
|
13 |
by (rtac (major RS disjE) 1);
|
|
14 |
by (rtac notE 1);
|
|
15 |
by (etac minor 2);
|
|
16 |
by (ALLGOALS assume_tac);
|
|
17 |
qed "make_neg_rule";
|
|
18 |
|
|
19 |
(*For Plaisted's "Postive refinement" of the MESON procedure*)
|
|
20 |
Goal "~P|Q ==> (P ==> Q)";
|
|
21 |
by (Blast_tac 1);
|
|
22 |
qed "make_refined_neg_rule";
|
|
23 |
|
|
24 |
(*P should be a literal*)
|
|
25 |
val [major,minor] = Goal "P|Q ==> ((P==>~P) ==> Q)";
|
|
26 |
by (rtac (major RS disjE) 1);
|
|
27 |
by (rtac notE 1);
|
|
28 |
by (etac minor 1);
|
|
29 |
by (ALLGOALS assume_tac);
|
|
30 |
qed "make_pos_rule";
|
|
31 |
|
|
32 |
|
|
33 |
(* Generation of a goal clause -- put away the final literal *)
|
|
34 |
|
|
35 |
val [major,minor] = Goal "~P ==> ((~P==>P) ==> False)";
|
|
36 |
by (rtac notE 1);
|
|
37 |
by (rtac minor 2);
|
|
38 |
by (ALLGOALS (rtac major));
|
|
39 |
qed "make_neg_goal";
|
|
40 |
|
|
41 |
val [major,minor] = Goal "P ==> ((P==>~P) ==> False)";
|
|
42 |
by (rtac notE 1);
|
|
43 |
by (rtac minor 1);
|
|
44 |
by (ALLGOALS (rtac major));
|
|
45 |
qed "make_pos_goal";
|
|
46 |
|
|
47 |
|
|
48 |
(* Lemmas for forward proof (like congruence rules) *)
|
|
49 |
|
|
50 |
(*NOTE: could handle conjunctions (faster?) by
|
|
51 |
nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *)
|
|
52 |
val major::prems = Goal
|
|
53 |
"[| P'&Q'; P' ==> P; Q' ==> Q |] ==> P&Q";
|
|
54 |
by (rtac (major RS conjE) 1);
|
|
55 |
by (rtac conjI 1);
|
|
56 |
by (ALLGOALS (eresolve_tac prems));
|
|
57 |
qed "conj_forward";
|
|
58 |
|
|
59 |
val major::prems = Goal
|
|
60 |
"[| P'|Q'; P' ==> P; Q' ==> Q |] ==> P|Q";
|
|
61 |
by (rtac (major RS disjE) 1);
|
|
62 |
by (ALLGOALS (dresolve_tac prems));
|
|
63 |
by (ALLGOALS (eresolve_tac [disjI1,disjI2]));
|
|
64 |
qed "disj_forward";
|
|
65 |
|
|
66 |
(*Version for removal of duplicate literals*)
|
|
67 |
val major::prems = Goal
|
|
68 |
"[| P'|Q'; P' ==> P; [| Q'; P==>False |] ==> Q |] ==> P|Q";
|
|
69 |
by (cut_facts_tac [major] 1);
|
|
70 |
by (blast_tac (claset() addIs prems) 1);
|
|
71 |
qed "disj_forward2";
|
|
72 |
|
|
73 |
val major::prems = Goal
|
|
74 |
"[| ALL x. P'(x); !!x. P'(x) ==> P(x) |] ==> ALL x. P(x)";
|
|
75 |
by (rtac allI 1);
|
|
76 |
by (resolve_tac prems 1);
|
|
77 |
by (rtac (major RS spec) 1);
|
|
78 |
qed "all_forward";
|
|
79 |
|
|
80 |
val major::prems = Goal
|
|
81 |
"[| EX x. P'(x); !!x. P'(x) ==> P(x) |] ==> EX x. P(x)";
|
|
82 |
by (rtac (major RS exE) 1);
|
|
83 |
by (rtac exI 1);
|
|
84 |
by (eresolve_tac prems 1);
|
|
85 |
qed "ex_forward";
|