| author | wenzelm | 
| Wed, 09 May 2007 19:37:21 +0200 | |
| changeset 22896 | 1c2abcabea61 | 
| parent 22846 | fb79144af9a3 | 
| permissions | -rw-r--r-- | 
| 
19159
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
1  | 
(* ID: $Id$  | 
| 20773 | 2  | 
Author: Jia Meng, NICTA  | 
| 
19159
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
3  | 
|
| 20773 | 4  | 
ATP rules.  | 
| 
19159
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
6  | 
|
| 19170 | 7  | 
signature RES_ATPSET =  | 
8  | 
sig  | 
|
| 21506 | 9  | 
val print_atpset: Proof.context -> unit  | 
10  | 
val get_atpset: Proof.context -> thm list  | 
|
| 20773 | 11  | 
val atp_add: attribute  | 
12  | 
val atp_del: attribute  | 
|
13  | 
val setup: theory -> theory  | 
|
| 19170 | 14  | 
end;  | 
15  | 
||
| 20773 | 16  | 
structure ResAtpset: RES_ATPSET =  | 
| 
19159
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
17  | 
struct  | 
| 
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
18  | 
|
| 20773 | 19  | 
structure Data = GenericDataFun  | 
20  | 
(  | 
|
21  | 
type T = thm list;  | 
|
22  | 
val empty = [];  | 
|
23  | 
val extend = I;  | 
|
24  | 
fun merge _ = Drule.merge_rules;  | 
|
25  | 
);  | 
|
| 
19159
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
26  | 
|
| 21506 | 27  | 
val get_atpset = Data.get o Context.Proof;  | 
| 
19159
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
28  | 
|
| 22846 | 29  | 
fun print_atpset ctxt =  | 
30  | 
Pretty.writeln (Pretty.big_list "ATP rules:"  | 
|
31  | 
(map (ProofContext.pretty_thm ctxt) (get_atpset ctxt)));  | 
|
32  | 
||
| 20773 | 33  | 
val atp_add = Thm.declaration_attribute (Data.map o Drule.add_rule);  | 
34  | 
val atp_del = Thm.declaration_attribute (Data.map o Drule.del_rule);  | 
|
| 
19159
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
35  | 
|
| 20773 | 36  | 
val setup =  | 
37  | 
  Attrib.add_attributes [("atp", Attrib.add_del_args atp_add atp_del, "declaration of ATP rules")];
 | 
|
| 
19159
 
f737ecbad1c4
A new file that sets up rules set used by ATPs. Rules are added to and removed from the set using "atp" attribute.
 
mengj 
parents:  
diff
changeset
 | 
38  | 
|
| 20547 | 39  | 
end;  |