| author | paulson | 
| Mon, 02 Aug 1999 11:24:01 +0200 | |
| changeset 7142 | 89e0ff71d113 | 
| parent 7127 | 48e235179ffb | 
| child 7213 | 08a354bbc34c | 
| permissions | -rw-r--r-- | 
| 1465 | 1  | 
(* Title: HOL/simpdata.ML  | 
| 923 | 2  | 
ID: $Id$  | 
| 1465 | 3  | 
Author: Tobias Nipkow  | 
| 923 | 4  | 
Copyright 1991 University of Cambridge  | 
5  | 
||
| 5304 | 6  | 
Instantiation of the generic simplifier for HOL.  | 
| 923 | 7  | 
*)  | 
8  | 
||
| 1984 | 9  | 
section "Simplifier";  | 
10  | 
||
| 6514 | 11  | 
(*** Addition of rules to simpsets and clasets simultaneously ***) (* FIXME move to Provers/clasimp.ML? *)  | 
| 1984 | 12  | 
|
| 
5190
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
13  | 
infix 4 addIffs delIffs;  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
14  | 
|
| 1984 | 15  | 
(*Takes UNCONDITIONAL theorems of the form A<->B to  | 
| 2031 | 16  | 
the Safe Intr rule B==>A and  | 
17  | 
the Safe Destruct rule A==>B.  | 
|
| 1984 | 18  | 
Also ~A goes to the Safe Elim rule A ==> ?R  | 
19  | 
Failing other cases, A is added as a Safe Intr rule*)  | 
|
20  | 
local  | 
|
21  | 
val iff_const = HOLogic.eq_const HOLogic.boolT;  | 
|
22  | 
||
| 
5190
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
23  | 
fun addIff ((cla, simp), th) =  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
24  | 
(case HOLogic.dest_Trueprop (#prop (rep_thm th)) of  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
25  | 
                (Const("Not", _) $ A) =>
 | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
26  | 
cla addSEs [zero_var_indexes (th RS notE)]  | 
| 2031 | 27  | 
| (con $ _ $ _) =>  | 
| 
5190
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
28  | 
if con = iff_const  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
29  | 
then cla addSIs [zero_var_indexes (th RS iffD2)]  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
30  | 
addSDs [zero_var_indexes (th RS iffD1)]  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
31  | 
else cla addSIs [th]  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
32  | 
| _ => cla addSIs [th],  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
33  | 
simp addsimps [th])  | 
| 6968 | 34  | 
      handle TERM _ => error ("AddIffs: theorem must be unconditional\n" ^ 
 | 
| 
5190
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
35  | 
string_of_thm th);  | 
| 1984 | 36  | 
|
| 
5190
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
37  | 
fun delIff ((cla, simp), th) =  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
38  | 
(case HOLogic.dest_Trueprop (#prop (rep_thm th)) of  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
39  | 
                (Const ("Not", _) $ A) =>
 | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
40  | 
cla delrules [zero_var_indexes (th RS notE)]  | 
| 2031 | 41  | 
| (con $ _ $ _) =>  | 
| 
5190
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
42  | 
if con = iff_const  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
43  | 
then cla delrules [zero_var_indexes (th RS iffD2),  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
44  | 
make_elim (zero_var_indexes (th RS iffD1))]  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
45  | 
else cla delrules [th]  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
46  | 
| _ => cla delrules [th],  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
47  | 
simp delsimps [th])  | 
| 6968 | 48  | 
      handle TERM _ => (warning("DelIffs: ignoring conditional theorem\n" ^ 
 | 
| 
5190
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
49  | 
string_of_thm th); (cla, simp));  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
50  | 
|
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
51  | 
fun store_clasimp (cla, simp) = (claset_ref () := cla; simpset_ref () := simp)  | 
| 1984 | 52  | 
in  | 
| 
5190
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
53  | 
val op addIffs = foldl addIff;  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
54  | 
val op delIffs = foldl delIff;  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
55  | 
fun AddIffs thms = store_clasimp ((claset (), simpset ()) addIffs thms);  | 
| 
 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 
berghofe 
parents: 
4930 
diff
changeset
 | 
56  | 
fun DelIffs thms = store_clasimp ((claset (), simpset ()) delIffs thms);  | 
| 1984 | 57  | 
end;  | 
58  | 
||
| 5304 | 59  | 
|
| 6514 | 60  | 
(* "iff" attribute *)  | 
61  | 
||
62  | 
local  | 
|
63  | 
fun change_global_css f (thy, th) =  | 
|
64  | 
let  | 
|
65  | 
val cs_ref = Classical.claset_ref_of thy;  | 
|
66  | 
val ss_ref = Simplifier.simpset_ref_of thy;  | 
|
67  | 
val (cs', ss') = f ((! cs_ref, ! ss_ref), [th]);  | 
|
68  | 
in cs_ref := cs'; ss_ref := ss'; (thy, th) end;  | 
|
69  | 
||
70  | 
fun change_local_css f (ctxt, th) =  | 
|
71  | 
let  | 
|
72  | 
val cs = Classical.get_local_claset ctxt;  | 
|
73  | 
val ss = Simplifier.get_local_simpset ctxt;  | 
|
74  | 
val (cs', ss') = f ((cs, ss), [th]);  | 
|
75  | 
val ctxt' =  | 
|
76  | 
ctxt  | 
|
77  | 
|> Classical.put_local_claset cs'  | 
|
78  | 
|> Simplifier.put_local_simpset ss';  | 
|
79  | 
in (ctxt', th) end;  | 
|
80  | 
in  | 
|
81  | 
||
82  | 
val iff_add_global = change_global_css (op addIffs);  | 
|
83  | 
val iff_add_local = change_local_css (op addIffs);  | 
|
84  | 
||
85  | 
val simpdata_setup =  | 
|
86  | 
  [Attrib.add_attributes [("iff", (Attrib.no_args iff_add_global, Attrib.no_args iff_add_local),
 | 
|
87  | 
"add rules to simpset and claset simultaneously")]];  | 
|
88  | 
||
89  | 
end;  | 
|
90  | 
||
91  | 
||
| 7031 | 92  | 
val [prem] = goal HOL.thy "x==y ==> x=y";  | 
93  | 
by (rewtac prem);  | 
|
94  | 
by (rtac refl 1);  | 
|
95  | 
qed "meta_eq_to_obj_eq";  | 
|
| 4640 | 96  | 
|
| 923 | 97  | 
local  | 
98  | 
||
| 7031 | 99  | 
fun prover s = prove_goal HOL.thy s (fn _ => [(Blast_tac 1)]);  | 
| 923 | 100  | 
|
| 2134 | 101  | 
in  | 
102  | 
||
| 5552 | 103  | 
(*Make meta-equalities. The operator below is Trueprop*)  | 
104  | 
||
| 6128 | 105  | 
fun mk_meta_eq r = r RS eq_reflection;  | 
106  | 
||
107  | 
val Eq_TrueI = mk_meta_eq(prover "P --> (P = True)" RS mp);  | 
|
108  | 
val Eq_FalseI = mk_meta_eq(prover "~P --> (P = False)" RS mp);  | 
|
| 5304 | 109  | 
|
| 6128 | 110  | 
fun mk_eq th = case concl_of th of  | 
111  | 
        Const("==",_)$_$_       => th
 | 
|
112  | 
    |   _$(Const("op =",_)$_$_) => mk_meta_eq th
 | 
|
113  | 
    |   _$(Const("Not",_)$_)    => th RS Eq_FalseI
 | 
|
114  | 
| _ => th RS Eq_TrueI;  | 
|
115  | 
(* last 2 lines requires all formulae to be of the from Trueprop(.) *)  | 
|
| 5304 | 116  | 
|
| 6128 | 117  | 
fun mk_eq_True r = Some(r RS meta_eq_to_obj_eq RS Eq_TrueI);  | 
| 5552 | 118  | 
|
| 6128 | 119  | 
fun mk_meta_cong rl =  | 
120  | 
standard(mk_meta_eq(replicate (nprems_of rl) meta_eq_to_obj_eq MRS rl))  | 
|
121  | 
handle THM _ =>  | 
|
122  | 
  error("Premises and conclusion of congruence rules must be =-equalities");
 | 
|
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3842 
diff
changeset
 | 
123  | 
|
| 
5975
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
124  | 
val not_not = prover "(~ ~ P) = P";  | 
| 923 | 125  | 
|
| 
5975
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
126  | 
val simp_thms = [not_not] @ map prover  | 
| 2082 | 127  | 
[ "(x=x) = True",  | 
| 
5975
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
128  | 
"(~True) = False", "(~False) = True",  | 
| 2082 | 129  | 
"(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))",  | 
| 4640 | 130  | 
"(True=P) = P", "(P=True) = P", "(False=P) = (~P)", "(P=False) = (~P)",  | 
| 2082 | 131  | 
"(True --> P) = P", "(False --> P) = True",  | 
132  | 
"(P --> True) = True", "(P --> P) = True",  | 
|
133  | 
"(P --> False) = (~P)", "(P --> ~P) = (~P)",  | 
|
134  | 
"(P & True) = P", "(True & P) = P",  | 
|
| 2800 | 135  | 
"(P & False) = False", "(False & P) = False",  | 
136  | 
"(P & P) = P", "(P & (P & Q)) = (P & Q)",  | 
|
| 3913 | 137  | 
"(P & ~P) = False", "(~P & P) = False",  | 
| 2082 | 138  | 
"(P | True) = True", "(True | P) = True",  | 
| 2800 | 139  | 
"(P | False) = P", "(False | P) = P",  | 
140  | 
"(P | P) = P", "(P | (P | Q)) = (P | Q)",  | 
|
| 3913 | 141  | 
"(P | ~P) = True", "(~P | P) = True",  | 
| 2082 | 142  | 
"((~P) = (~Q)) = (P=Q)",  | 
| 3842 | 143  | 
"(!x. P) = P", "(? x. P) = P", "? x. x=t", "? x. t=x",  | 
| 4351 | 144  | 
(*two needed for the one-point-rule quantifier simplification procs*)  | 
145  | 
"(? x. x=t & P(x)) = P(t)", (*essential for termination!!*)  | 
|
146  | 
"(! x. t=x --> P(x)) = P(t)" ]; (*covers a stray case*)  | 
|
| 923 | 147  | 
|
| 5552 | 148  | 
(* Add congruence rules for = (instead of ==) *)  | 
| 4351 | 149  | 
|
| 5552 | 150  | 
(* ###FIXME: Move to simplifier,  | 
151  | 
taking mk_meta_cong as input, eliminating addeqcongs and deleqcongs *)  | 
|
152  | 
infix 4 addcongs delcongs;  | 
|
| 4640 | 153  | 
fun ss addcongs congs = ss addeqcongs (map mk_meta_cong congs);  | 
154  | 
fun ss delcongs congs = ss deleqcongs (map mk_meta_cong congs);  | 
|
| 4086 | 155  | 
fun Addcongs congs = (simpset_ref() := simpset() addcongs congs);  | 
156  | 
fun Delcongs congs = (simpset_ref() := simpset() delcongs congs);  | 
|
| 1264 | 157  | 
|
| 5552 | 158  | 
|
| 1922 | 159  | 
val imp_cong = impI RSN  | 
160  | 
(2, prove_goal HOL.thy "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))"  | 
|
| 7031 | 161  | 
(fn _=> [(Blast_tac 1)]) RS mp RS mp);  | 
| 1922 | 162  | 
|
| 
1948
 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 
paulson 
parents: 
1922 
diff
changeset
 | 
163  | 
(*Miniscoping: pushing in existential quantifiers*)  | 
| 
 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 
paulson 
parents: 
1922 
diff
changeset
 | 
164  | 
val ex_simps = map prover  | 
| 3842 | 165  | 
["(EX x. P x & Q) = ((EX x. P x) & Q)",  | 
166  | 
"(EX x. P & Q x) = (P & (EX x. Q x))",  | 
|
167  | 
"(EX x. P x | Q) = ((EX x. P x) | Q)",  | 
|
168  | 
"(EX x. P | Q x) = (P | (EX x. Q x))",  | 
|
169  | 
"(EX x. P x --> Q) = ((ALL x. P x) --> Q)",  | 
|
170  | 
"(EX x. P --> Q x) = (P --> (EX x. Q x))"];  | 
|
| 
1948
 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 
paulson 
parents: 
1922 
diff
changeset
 | 
171  | 
|
| 
 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 
paulson 
parents: 
1922 
diff
changeset
 | 
172  | 
(*Miniscoping: pushing in universal quantifiers*)  | 
| 
 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 
paulson 
parents: 
1922 
diff
changeset
 | 
173  | 
val all_simps = map prover  | 
| 3842 | 174  | 
["(ALL x. P x & Q) = ((ALL x. P x) & Q)",  | 
175  | 
"(ALL x. P & Q x) = (P & (ALL x. Q x))",  | 
|
176  | 
"(ALL x. P x | Q) = ((ALL x. P x) | Q)",  | 
|
177  | 
"(ALL x. P | Q x) = (P | (ALL x. Q x))",  | 
|
178  | 
"(ALL x. P x --> Q) = ((EX x. P x) --> Q)",  | 
|
179  | 
"(ALL x. P --> Q x) = (P --> (ALL x. Q x))"];  | 
|
| 
1948
 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 
paulson 
parents: 
1922 
diff
changeset
 | 
180  | 
|
| 923 | 181  | 
|
| 2022 | 182  | 
(* elimination of existential quantifiers in assumptions *)  | 
| 923 | 183  | 
|
184  | 
val ex_all_equiv =  | 
|
185  | 
let val lemma1 = prove_goal HOL.thy  | 
|
186  | 
"(? x. P(x) ==> PROP Q) ==> (!!x. P(x) ==> PROP Q)"  | 
|
187  | 
(fn prems => [resolve_tac prems 1, etac exI 1]);  | 
|
188  | 
val lemma2 = prove_goalw HOL.thy [Ex_def]  | 
|
189  | 
"(!!x. P(x) ==> PROP Q) ==> (? x. P(x) ==> PROP Q)"  | 
|
| 7031 | 190  | 
(fn prems => [(REPEAT(resolve_tac prems 1))])  | 
| 923 | 191  | 
in equal_intr lemma1 lemma2 end;  | 
192  | 
||
193  | 
end;  | 
|
194  | 
||
| 3654 | 195  | 
(* Elimination of True from asumptions: *)  | 
196  | 
||
197  | 
val True_implies_equals = prove_goal HOL.thy  | 
|
198  | 
"(True ==> PROP P) == PROP P"  | 
|
| 7031 | 199  | 
(fn _ => [rtac equal_intr_rule 1, atac 2,  | 
| 3654 | 200  | 
METAHYPS (fn prems => resolve_tac prems 1) 1,  | 
201  | 
rtac TrueI 1]);  | 
|
202  | 
||
| 7031 | 203  | 
fun prove nm thm = qed_goal nm HOL.thy thm (fn _ => [(Blast_tac 1)]);  | 
| 923 | 204  | 
|
205  | 
prove "conj_commute" "(P&Q) = (Q&P)";  | 
|
206  | 
prove "conj_left_commute" "(P&(Q&R)) = (Q&(P&R))";  | 
|
207  | 
val conj_comms = [conj_commute, conj_left_commute];  | 
|
| 2134 | 208  | 
prove "conj_assoc" "((P&Q)&R) = (P&(Q&R))";  | 
| 923 | 209  | 
|
| 1922 | 210  | 
prove "disj_commute" "(P|Q) = (Q|P)";  | 
211  | 
prove "disj_left_commute" "(P|(Q|R)) = (Q|(P|R))";  | 
|
212  | 
val disj_comms = [disj_commute, disj_left_commute];  | 
|
| 2134 | 213  | 
prove "disj_assoc" "((P|Q)|R) = (P|(Q|R))";  | 
| 1922 | 214  | 
|
| 923 | 215  | 
prove "conj_disj_distribL" "(P&(Q|R)) = (P&Q | P&R)";  | 
216  | 
prove "conj_disj_distribR" "((P|Q)&R) = (P&R | Q&R)";  | 
|
| 
1485
 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 
nipkow 
parents: 
1465 
diff
changeset
 | 
217  | 
|
| 1892 | 218  | 
prove "disj_conj_distribL" "(P|(Q&R)) = ((P|Q) & (P|R))";  | 
219  | 
prove "disj_conj_distribR" "((P&Q)|R) = ((P|R) & (Q|R))";  | 
|
220  | 
||
| 2134 | 221  | 
prove "imp_conjR" "(P --> (Q&R)) = ((P-->Q) & (P-->R))";  | 
222  | 
prove "imp_conjL" "((P&Q) -->R) = (P --> (Q --> R))";  | 
|
223  | 
prove "imp_disjL" "((P|Q) --> R) = ((P-->R)&(Q-->R))";  | 
|
| 1892 | 224  | 
|
| 3448 | 225  | 
(*These two are specialized, but imp_disj_not1 is useful in Auth/Yahalom.ML*)  | 
226  | 
prove "imp_disj_not1" "((P --> Q | R)) = (~Q --> P --> R)";  | 
|
227  | 
prove "imp_disj_not2" "((P --> Q | R)) = (~R --> P --> Q)";  | 
|
228  | 
||
| 3904 | 229  | 
prove "imp_disj1" "((P-->Q)|R) = (P--> Q|R)";  | 
230  | 
prove "imp_disj2" "(Q|(P-->R)) = (P--> Q|R)";  | 
|
231  | 
||
| 
1485
 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 
nipkow 
parents: 
1465 
diff
changeset
 | 
232  | 
prove "de_Morgan_disj" "(~(P | Q)) = (~P & ~Q)";  | 
| 
 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 
nipkow 
parents: 
1465 
diff
changeset
 | 
233  | 
prove "de_Morgan_conj" "(~(P & Q)) = (~P | ~Q)";  | 
| 
3446
 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 
paulson 
parents: 
3282 
diff
changeset
 | 
234  | 
prove "not_imp" "(~(P --> Q)) = (P & ~Q)";  | 
| 1922 | 235  | 
prove "not_iff" "(P~=Q) = (P = (~Q))";  | 
| 4743 | 236  | 
prove "disj_not1" "(~P | Q) = (P --> Q)";  | 
237  | 
prove "disj_not2" "(P | ~Q) = (Q --> P)"; (* changes orientation :-( *)  | 
|
| 
5975
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
238  | 
prove "imp_conv_disj" "(P --> Q) = ((~P) | Q)";  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
239  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
240  | 
prove "iff_conv_conj_imp" "(P = Q) = ((P --> Q) & (Q --> P))";  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
241  | 
|
| 
1485
 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 
nipkow 
parents: 
1465 
diff
changeset
 | 
242  | 
|
| 4830 | 243  | 
(*Avoids duplication of subgoals after split_if, when the true and false  | 
| 2134 | 244  | 
cases boil down to the same thing.*)  | 
245  | 
prove "cases_simp" "((P --> Q) & (~P --> Q)) = Q";  | 
|
246  | 
||
| 3842 | 247  | 
prove "not_all" "(~ (! x. P(x))) = (? x.~P(x))";  | 
| 1922 | 248  | 
prove "imp_all" "((! x. P x) --> Q) = (? x. P x --> Q)";  | 
| 3842 | 249  | 
prove "not_ex" "(~ (? x. P(x))) = (! x.~P(x))";  | 
| 1922 | 250  | 
prove "imp_ex" "((? x. P x) --> Q) = (! x. P x --> Q)";  | 
| 1660 | 251  | 
|
| 1655 | 252  | 
prove "ex_disj_distrib" "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))";  | 
253  | 
prove "all_conj_distrib" "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))";  | 
|
254  | 
||
| 2134 | 255  | 
(* '&' congruence rule: not included by default!  | 
256  | 
May slow rewrite proofs down by as much as 50% *)  | 
|
257  | 
||
258  | 
let val th = prove_goal HOL.thy  | 
|
259  | 
"(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))"  | 
|
| 7031 | 260  | 
(fn _=> [(Blast_tac 1)])  | 
| 2134 | 261  | 
in  bind_thm("conj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
 | 
262  | 
||
263  | 
let val th = prove_goal HOL.thy  | 
|
264  | 
"(Q=Q')--> (Q'--> (P=P'))--> ((P&Q) = (P'&Q'))"  | 
|
| 7031 | 265  | 
(fn _=> [(Blast_tac 1)])  | 
| 2134 | 266  | 
in  bind_thm("rev_conj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
 | 
267  | 
||
268  | 
(* '|' congruence rule: not included by default! *)  | 
|
269  | 
||
270  | 
let val th = prove_goal HOL.thy  | 
|
271  | 
"(P=P')--> (~P'--> (Q=Q'))--> ((P|Q) = (P'|Q'))"  | 
|
| 7031 | 272  | 
(fn _=> [(Blast_tac 1)])  | 
| 2134 | 273  | 
in  bind_thm("disj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
 | 
274  | 
||
275  | 
prove "eq_sym_conv" "(x=y) = (y=x)";  | 
|
276  | 
||
| 5278 | 277  | 
|
278  | 
(** if-then-else rules **)  | 
|
279  | 
||
| 7031 | 280  | 
Goalw [if_def] "(if True then x else y) = x";  | 
281  | 
by (Blast_tac 1);  | 
|
282  | 
qed "if_True";  | 
|
| 2134 | 283  | 
|
| 7031 | 284  | 
Goalw [if_def] "(if False then x else y) = y";  | 
285  | 
by (Blast_tac 1);  | 
|
286  | 
qed "if_False";  | 
|
| 2134 | 287  | 
|
| 
7127
 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 
paulson 
parents: 
7031 
diff
changeset
 | 
288  | 
Goalw [if_def] "P ==> (if P then x else y) = x";  | 
| 7031 | 289  | 
by (Blast_tac 1);  | 
290  | 
qed "if_P";  | 
|
| 5304 | 291  | 
|
| 
7127
 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 
paulson 
parents: 
7031 
diff
changeset
 | 
292  | 
Goalw [if_def] "~P ==> (if P then x else y) = y";  | 
| 7031 | 293  | 
by (Blast_tac 1);  | 
294  | 
qed "if_not_P";  | 
|
| 2134 | 295  | 
|
| 7031 | 296  | 
Goal "P(if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))";  | 
297  | 
by (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1);
 | 
|
298  | 
by (stac if_P 2);  | 
|
299  | 
by (stac if_not_P 1);  | 
|
300  | 
by (ALLGOALS (Blast_tac));  | 
|
301  | 
qed "split_if";  | 
|
302  | 
||
| 4830 | 303  | 
(* for backwards compatibility: *)  | 
304  | 
val expand_if = split_if;  | 
|
| 
4205
 
96632970d203
simpdata.ML: renamed split_prem_tac to split_asm_tac, added split_if_asm
 
oheimb 
parents: 
4189 
diff
changeset
 | 
305  | 
|
| 7031 | 306  | 
Goal "P(if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))";  | 
307  | 
by (stac split_if 1);  | 
|
308  | 
by (Blast_tac 1);  | 
|
309  | 
qed "split_if_asm";  | 
|
| 2134 | 310  | 
|
| 7031 | 311  | 
Goal "(if c then x else x) = x";  | 
312  | 
by (stac split_if 1);  | 
|
313  | 
by (Blast_tac 1);  | 
|
314  | 
qed "if_cancel";  | 
|
| 5304 | 315  | 
|
| 7031 | 316  | 
Goal "(if x = y then y else x) = x";  | 
317  | 
by (stac split_if 1);  | 
|
318  | 
by (Blast_tac 1);  | 
|
319  | 
qed "if_eq_cancel";  | 
|
| 5304 | 320  | 
|
| 
4769
 
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
 
paulson 
parents: 
4744 
diff
changeset
 | 
321  | 
(*This form is useful for expanding IFs on the RIGHT of the ==> symbol*)  | 
| 
7127
 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 
paulson 
parents: 
7031 
diff
changeset
 | 
322  | 
Goal "(if P then Q else R) = ((P-->Q) & (~P-->R))";  | 
| 7031 | 323  | 
by (rtac split_if 1);  | 
324  | 
qed "if_bool_eq_conj";  | 
|
| 
4769
 
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
 
paulson 
parents: 
4744 
diff
changeset
 | 
325  | 
|
| 
 
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
 
paulson 
parents: 
4744 
diff
changeset
 | 
326  | 
(*And this form is useful for expanding IFs on the LEFT*)  | 
| 7031 | 327  | 
Goal "(if P then Q else R) = ((P&Q) | (~P&R))";  | 
328  | 
by (stac split_if 1);  | 
|
329  | 
by (Blast_tac 1);  | 
|
330  | 
qed "if_bool_eq_disj";  | 
|
| 2134 | 331  | 
|
| 4351 | 332  | 
|
333  | 
(*** make simplification procedures for quantifier elimination ***)  | 
|
334  | 
||
335  | 
structure Quantifier1 = Quantifier1Fun(  | 
|
336  | 
struct  | 
|
337  | 
(*abstract syntax*)  | 
|
338  | 
  fun dest_eq((c as Const("op =",_)) $ s $ t) = Some(c,s,t)
 | 
|
339  | 
| dest_eq _ = None;  | 
|
340  | 
  fun dest_conj((c as Const("op &",_)) $ s $ t) = Some(c,s,t)
 | 
|
341  | 
| dest_conj _ = None;  | 
|
342  | 
val conj = HOLogic.conj  | 
|
343  | 
val imp = HOLogic.imp  | 
|
344  | 
(*rules*)  | 
|
345  | 
val iff_reflection = eq_reflection  | 
|
346  | 
val iffI = iffI  | 
|
347  | 
val sym = sym  | 
|
348  | 
val conjI= conjI  | 
|
349  | 
val conjE= conjE  | 
|
350  | 
val impI = impI  | 
|
351  | 
val impE = impE  | 
|
352  | 
val mp = mp  | 
|
353  | 
val exI = exI  | 
|
354  | 
val exE = exE  | 
|
355  | 
val allI = allI  | 
|
356  | 
val allE = allE  | 
|
357  | 
end);  | 
|
358  | 
||
| 
4320
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
359  | 
local  | 
| 
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
360  | 
val ex_pattern =  | 
| 6394 | 361  | 
  Thm.read_cterm (Theory.sign_of HOL.thy) ("EX x. P(x) & Q(x)",HOLogic.boolT)
 | 
| 3913 | 362  | 
|
| 
4320
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
363  | 
val all_pattern =  | 
| 6394 | 364  | 
  Thm.read_cterm (Theory.sign_of HOL.thy) ("ALL x. P(x) & P'(x) --> Q(x)",HOLogic.boolT)
 | 
| 
4320
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
365  | 
|
| 
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
366  | 
in  | 
| 
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
367  | 
val defEX_regroup =  | 
| 
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
368  | 
mk_simproc "defined EX" [ex_pattern] Quantifier1.rearrange_ex;  | 
| 
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
369  | 
val defALL_regroup =  | 
| 
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
370  | 
mk_simproc "defined ALL" [all_pattern] Quantifier1.rearrange_all;  | 
| 
 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 
nipkow 
parents: 
4205 
diff
changeset
 | 
371  | 
end;  | 
| 3913 | 372  | 
|
| 4351 | 373  | 
|
374  | 
(*** Case splitting ***)  | 
|
| 3913 | 375  | 
|
| 5304 | 376  | 
structure SplitterData =  | 
377  | 
struct  | 
|
378  | 
structure Simplifier = Simplifier  | 
|
| 5552 | 379  | 
val mk_eq = mk_eq  | 
| 5304 | 380  | 
val meta_eq_to_iff = meta_eq_to_obj_eq  | 
381  | 
val iffD = iffD2  | 
|
382  | 
val disjE = disjE  | 
|
383  | 
val conjE = conjE  | 
|
384  | 
val exE = exE  | 
|
385  | 
val contrapos = contrapos  | 
|
386  | 
val contrapos2 = contrapos2  | 
|
387  | 
val notnotD = notnotD  | 
|
388  | 
end;  | 
|
| 4681 | 389  | 
|
| 5304 | 390  | 
structure Splitter = SplitterFun(SplitterData);  | 
| 2263 | 391  | 
|
| 5304 | 392  | 
val split_tac = Splitter.split_tac;  | 
393  | 
val split_inside_tac = Splitter.split_inside_tac;  | 
|
394  | 
val split_asm_tac = Splitter.split_asm_tac;  | 
|
| 5307 | 395  | 
val op addsplits = Splitter.addsplits;  | 
396  | 
val op delsplits = Splitter.delsplits;  | 
|
| 5304 | 397  | 
val Addsplits = Splitter.Addsplits;  | 
398  | 
val Delsplits = Splitter.Delsplits;  | 
|
| 
4718
 
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
 
oheimb 
parents: 
4681 
diff
changeset
 | 
399  | 
|
| 2134 | 400  | 
(*In general it seems wrong to add distributive laws by default: they  | 
401  | 
might cause exponential blow-up. But imp_disjL has been in for a while  | 
|
402  | 
and cannot be removed without affecting existing proofs. Moreover,  | 
|
403  | 
rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the  | 
|
404  | 
grounds that it allows simplification of R in the two cases.*)  | 
|
405  | 
||
| 5304 | 406  | 
fun gen_all th = forall_elim_vars (#maxidx(rep_thm th)+1) th;  | 
407  | 
||
| 2134 | 408  | 
val mksimps_pairs =  | 
409  | 
  [("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
 | 
|
410  | 
   ("All", [spec]), ("True", []), ("False", []),
 | 
|
| 
4769
 
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
 
paulson 
parents: 
4744 
diff
changeset
 | 
411  | 
   ("If", [if_bool_eq_conj RS iffD1])];
 | 
| 1758 | 412  | 
|
| 5552 | 413  | 
(* ###FIXME: move to Provers/simplifier.ML  | 
| 5304 | 414  | 
val mk_atomize: (string * thm list) list -> thm -> thm list  | 
415  | 
*)  | 
|
| 5552 | 416  | 
(* ###FIXME: move to Provers/simplifier.ML *)  | 
| 5304 | 417  | 
fun mk_atomize pairs =  | 
418  | 
let fun atoms th =  | 
|
419  | 
(case concl_of th of  | 
|
420  | 
           Const("Trueprop",_) $ p =>
 | 
|
421  | 
(case head_of p of  | 
|
422  | 
Const(a,_) =>  | 
|
423  | 
(case assoc(pairs,a) of  | 
|
424  | 
Some(rls) => flat (map atoms ([th] RL rls))  | 
|
425  | 
| None => [th])  | 
|
426  | 
| _ => [th])  | 
|
427  | 
| _ => [th])  | 
|
428  | 
in atoms end;  | 
|
429  | 
||
| 5552 | 430  | 
fun mksimps pairs = (map mk_eq o mk_atomize pairs o gen_all);  | 
| 5304 | 431  | 
|
| 4640 | 432  | 
fun unsafe_solver prems = FIRST'[resolve_tac (reflexive_thm::TrueI::refl::prems),  | 
| 
2636
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
433  | 
atac, etac FalseE];  | 
| 
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
434  | 
(*No premature instantiation of variables during simplification*)  | 
| 4640 | 435  | 
fun safe_solver prems = FIRST'[match_tac (reflexive_thm::TrueI::prems),  | 
| 
2636
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
436  | 
eq_assume_tac, ematch_tac [FalseE]];  | 
| 
2443
 
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
 
oheimb 
parents: 
2263 
diff
changeset
 | 
437  | 
|
| 
2636
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
438  | 
val HOL_basic_ss = empty_ss setsubgoaler asm_simp_tac  | 
| 
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
439  | 
setSSolver safe_solver  | 
| 
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
440  | 
setSolver unsafe_solver  | 
| 4677 | 441  | 
setmksimps (mksimps mksimps_pairs)  | 
| 5552 | 442  | 
setmkeqTrue mk_eq_True;  | 
| 
2443
 
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
 
oheimb 
parents: 
2263 
diff
changeset
 | 
443  | 
|
| 
3446
 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 
paulson 
parents: 
3282 
diff
changeset
 | 
444  | 
val HOL_ss =  | 
| 
 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 
paulson 
parents: 
3282 
diff
changeset
 | 
445  | 
HOL_basic_ss addsimps  | 
| 
 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 
paulson 
parents: 
3282 
diff
changeset
 | 
446  | 
([triv_forall_equality, (* prunes params *)  | 
| 3654 | 447  | 
True_implies_equals, (* prune asms `True' *)  | 
| 
4718
 
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
 
oheimb 
parents: 
4681 
diff
changeset
 | 
448  | 
if_True, if_False, if_cancel, if_eq_cancel,  | 
| 5304 | 449  | 
imp_disjL, conj_assoc, disj_assoc,  | 
| 3904 | 450  | 
de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2, not_imp,  | 
| 
5447
 
df03d330aeab
Proved and added rewrite rule (@x. x=y) = y to simpset.
 
nipkow 
parents: 
5307 
diff
changeset
 | 
451  | 
disj_not1, not_all, not_ex, cases_simp, Eps_eq, Eps_sym_eq]  | 
| 
3446
 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 
paulson 
parents: 
3282 
diff
changeset
 | 
452  | 
@ ex_simps @ all_simps @ simp_thms)  | 
| 
4032
 
4b1c69d8b767
For each datatype `t' there is now a theorem `split_t_case' of the form
 
nipkow 
parents: 
3919 
diff
changeset
 | 
453  | 
addsimprocs [defALL_regroup,defEX_regroup]  | 
| 
4744
 
4469d498cd48
moved addsplits [expand_if] from HOL_basic_ss to HOL_ss;
 
wenzelm 
parents: 
4743 
diff
changeset
 | 
454  | 
addcongs [imp_cong]  | 
| 4830 | 455  | 
addsplits [split_if];  | 
| 2082 | 456  | 
|
| 6293 | 457  | 
(*Simplifies x assuming c and y assuming ~c*)  | 
458  | 
val prems = Goalw [if_def]  | 
|
459  | 
"[| b=c; c ==> x=u; ~c ==> y=v |] ==> \  | 
|
460  | 
\ (if b then x else y) = (if c then u else v)";  | 
|
461  | 
by (asm_simp_tac (HOL_ss addsimps prems) 1);  | 
|
462  | 
qed "if_cong";  | 
|
463  | 
||
| 
7127
 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 
paulson 
parents: 
7031 
diff
changeset
 | 
464  | 
(*Prevents simplification of x and y: faster and allows the execution  | 
| 
 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 
paulson 
parents: 
7031 
diff
changeset
 | 
465  | 
of functional programs. NOW THE DEFAULT.*)  | 
| 7031 | 466  | 
Goal "b=c ==> (if b then x else y) = (if c then x else y)";  | 
467  | 
by (etac arg_cong 1);  | 
|
468  | 
qed "if_weak_cong";  | 
|
| 6293 | 469  | 
|
470  | 
(*Prevents simplification of t: much faster*)  | 
|
| 7031 | 471  | 
Goal "a = b ==> (let x=a in t(x)) = (let x=b in t(x))";  | 
472  | 
by (etac arg_cong 1);  | 
|
473  | 
qed "let_weak_cong";  | 
|
| 6293 | 474  | 
|
| 7031 | 475  | 
Goal "f(if c then x else y) = (if c then f x else f y)";  | 
476  | 
by (simp_tac (HOL_ss setloop (split_tac [split_if])) 1);  | 
|
477  | 
qed "if_distrib";  | 
|
| 1655 | 478  | 
|
| 1984 | 479  | 
|
| 4327 | 480  | 
(*For expand_case_tac*)  | 
| 
2948
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
481  | 
val prems = goal HOL.thy "[| P ==> Q(True); ~P ==> Q(False) |] ==> Q(P)";  | 
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
482  | 
by (case_tac "P" 1);  | 
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
483  | 
by (ALLGOALS (asm_simp_tac (HOL_ss addsimps prems)));  | 
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
484  | 
val expand_case = result();  | 
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
485  | 
|
| 4327 | 486  | 
(*Used in Auth proofs. Typically P contains Vars that become instantiated  | 
487  | 
during unification.*)  | 
|
| 
2948
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
488  | 
fun expand_case_tac P i =  | 
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
489  | 
    res_inst_tac [("P",P)] expand_case i THEN
 | 
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
490  | 
Simp_tac (i+1) THEN  | 
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
491  | 
Simp_tac i;  | 
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
492  | 
|
| 
 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 
paulson 
parents: 
2935 
diff
changeset
 | 
493  | 
|
| 4119 | 494  | 
(* install implicit simpset *)  | 
| 1984 | 495  | 
|
| 6915 | 496  | 
simpset_ref() := HOL_ss addcongs [if_weak_cong];  | 
| 
3615
 
e5322197cfea
Moved some functions which used to be part of thy_data.ML
 
berghofe 
parents: 
3577 
diff
changeset
 | 
497  | 
|
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents: 
4651 
diff
changeset
 | 
498  | 
|
| 5219 | 499  | 
(*** integration of simplifier with classical reasoner ***)  | 
| 
2636
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
500  | 
|
| 5219 | 501  | 
structure Clasimp = ClasimpFun  | 
| 5552 | 502  | 
(structure Simplifier = Simplifier  | 
503  | 
and Classical = Classical  | 
|
504  | 
and Blast = Blast);  | 
|
| 
4652
 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 
oheimb 
parents: 
4651 
diff
changeset
 | 
505  | 
open Clasimp;  | 
| 
2636
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
506  | 
|
| 
 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 
oheimb 
parents: 
2595 
diff
changeset
 | 
507  | 
val HOL_css = (HOL_cs, HOL_ss);  | 
| 
5975
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
508  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
509  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
510  | 
(*** A general refutation procedure ***)  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
511  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
512  | 
(* Parameters:  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
513  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
514  | 
test: term -> bool  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
515  | 
tests if a term is at all relevant to the refutation proof;  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
516  | 
if not, then it can be discarded. Can improve performance,  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
517  | 
esp. if disjunctions can be discarded (no case distinction needed!).  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
518  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
519  | 
prep_tac: int -> tactic  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
520  | 
A preparation tactic to be applied to the goal once all relevant premises  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
521  | 
have been moved to the conclusion.  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
522  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
523  | 
ref_tac: int -> tactic  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
524  | 
the actual refutation tactic. Should be able to deal with goals  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
525  | 
[| A1; ...; An |] ==> False  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
526  | 
where the Ai are atomic, i.e. no top-level &, | or ?  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
527  | 
*)  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
528  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
529  | 
fun refute_tac test prep_tac ref_tac =  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
530  | 
let val nnf_simps =  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
531  | 
[imp_conv_disj,iff_conv_conj_imp,de_Morgan_disj,de_Morgan_conj,  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
532  | 
not_all,not_ex,not_not];  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
533  | 
val nnf_simpset =  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
534  | 
empty_ss setmkeqTrue mk_eq_True  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
535  | 
setmksimps (mksimps mksimps_pairs)  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
536  | 
addsimps nnf_simps;  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
537  | 
val prem_nnf_tac = full_simp_tac nnf_simpset;  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
538  | 
|
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
539  | 
val refute_prems_tac =  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
540  | 
REPEAT(eresolve_tac [conjE, exE] 1 ORELSE  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
541  | 
filter_prems_tac test 1 ORELSE  | 
| 6301 | 542  | 
etac disjE 1) THEN  | 
| 
5975
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
543  | 
ref_tac 1;  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
544  | 
in EVERY'[TRY o filter_prems_tac test,  | 
| 6128 | 545  | 
DETERM o REPEAT o etac rev_mp, prep_tac, rtac ccontr, prem_nnf_tac,  | 
| 
5975
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
546  | 
SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)]  | 
| 
 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 
nipkow 
parents: 
5552 
diff
changeset
 | 
547  | 
end;  |