| author | blanchet | 
| Tue, 09 Apr 2013 16:32:04 +0200 | |
| changeset 51652 | 5ff01d585a8c | 
| parent 51650 | 3dd495cd98a2 | 
| child 51998 | f732a674db1b | 
| permissions | -rw-r--r-- | 
| 46321 | 1  | 
(* Title: HOL/TPTP/atp_theory_export.ML  | 
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
3  | 
Copyright 2011  | 
| 
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
4  | 
|
| 48234 | 5  | 
Export Isabelle theories as first-order TPTP inferences.  | 
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
7  | 
|
| 46321 | 8  | 
signature ATP_THEORY_EXPORT =  | 
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
9  | 
sig  | 
| 45305 | 10  | 
type atp_format = ATP_Problem.atp_format  | 
11  | 
||
| 51646 | 12  | 
datatype inference_policy =  | 
13  | 
No_Inferences | Unchecked_Inferences | Checked_Inferences  | 
|
14  | 
||
| 48234 | 15  | 
val generate_atp_inference_file_for_theory :  | 
| 51646 | 16  | 
Proof.context -> theory -> atp_format -> inference_policy -> string  | 
17  | 
-> string -> unit  | 
|
18  | 
val generate_casc_lbt_isa_files_for_theory :  | 
|
19  | 
Proof.context -> theory -> atp_format -> inference_policy -> string  | 
|
20  | 
-> string -> unit  | 
|
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
21  | 
end;  | 
| 
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
22  | 
|
| 48234 | 23  | 
structure ATP_Theory_Export : ATP_THEORY_EXPORT =  | 
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
24  | 
struct  | 
| 
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
25  | 
|
| 
43479
 
5af1abc13c1f
only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
 
blanchet 
parents: 
43468 
diff
changeset
 | 
26  | 
open ATP_Problem  | 
| 
43566
 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 
blanchet 
parents: 
43499 
diff
changeset
 | 
27  | 
open ATP_Proof  | 
| 46320 | 28  | 
open ATP_Problem_Generate  | 
| 
43566
 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 
blanchet 
parents: 
43499 
diff
changeset
 | 
29  | 
open ATP_Systems  | 
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
30  | 
|
| 51646 | 31  | 
datatype inference_policy =  | 
32  | 
No_Inferences | Unchecked_Inferences | Checked_Inferences  | 
|
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
33  | 
|
| 51646 | 34  | 
val fact_name_of = prefix fact_prefix o ascii_of  | 
| 
43479
 
5af1abc13c1f
only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
 
blanchet 
parents: 
43468 
diff
changeset
 | 
35  | 
|
| 48318 | 36  | 
fun atp_for_format (THF (Polymorphic, _, _, _)) = dummy_thfN  | 
37  | 
| atp_for_format (THF (Monomorphic, _, _, _)) = satallaxN  | 
|
38  | 
| atp_for_format (DFG Polymorphic) = spass_polyN  | 
|
39  | 
| atp_for_format (DFG Monomorphic) = spassN  | 
|
40  | 
| atp_for_format (TFF (Polymorphic, _)) = alt_ergoN  | 
|
41  | 
| atp_for_format (TFF (Monomorphic, _)) = vampireN  | 
|
42  | 
| atp_for_format FOF = eN  | 
|
43  | 
| atp_for_format CNF_UEQ = waldmeisterN  | 
|
44  | 
| atp_for_format CNF = eN  | 
|
45  | 
||
| 51646 | 46  | 
val atp_timeout = seconds 0.5  | 
47  | 
||
48  | 
fun run_atp ctxt format problem =  | 
|
| 
43566
 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 
blanchet 
parents: 
43499 
diff
changeset
 | 
49  | 
let  | 
| 
 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 
blanchet 
parents: 
43499 
diff
changeset
 | 
50  | 
val thy = Proof_Context.theory_of ctxt  | 
| 
47055
 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 
blanchet 
parents: 
47038 
diff
changeset
 | 
51  | 
val prob_file = File.tmp_path (Path.explode "prob")  | 
| 48318 | 52  | 
val atp = atp_for_format format  | 
| 
47606
 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 
blanchet 
parents: 
47055 
diff
changeset
 | 
53  | 
    val {exec, arguments, proof_delims, known_failures, ...} =
 | 
| 
 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 
blanchet 
parents: 
47055 
diff
changeset
 | 
54  | 
get_atp thy atp ()  | 
| 
47038
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
46734 
diff
changeset
 | 
55  | 
val ord = effective_term_order ctxt atp  | 
| 
 
2409b484e1cc
continued implementation of term ordering attributes
 
blanchet 
parents: 
46734 
diff
changeset
 | 
56  | 
val _ = problem |> lines_for_atp_problem format ord (K [])  | 
| 
46442
 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 
blanchet 
parents: 
46409 
diff
changeset
 | 
57  | 
|> File.write_list prob_file  | 
| 
48376
 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 
blanchet 
parents: 
48324 
diff
changeset
 | 
58  | 
val path = getenv (List.last (fst exec)) ^ "/" ^ List.last (snd exec)  | 
| 
43566
 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 
blanchet 
parents: 
43499 
diff
changeset
 | 
59  | 
val command =  | 
| 50927 | 60  | 
File.shell_path (Path.explode path) ^ " " ^  | 
| 51646 | 61  | 
arguments ctxt false "" atp_timeout (File.shell_path prob_file)  | 
| 50927 | 62  | 
(ord, K [], K [])  | 
| 51646 | 63  | 
val outcome =  | 
64  | 
TimeLimit.timeLimit atp_timeout Isabelle_System.bash_output command  | 
|
65  | 
|> fst  | 
|
66  | 
|> extract_tstplike_proof_and_outcome false proof_delims known_failures  | 
|
67  | 
|> snd  | 
|
68  | 
handle TimeLimit.TimeOut => SOME TimedOut  | 
|
| 
51647
 
25acf689a53e
no need to filter tautologies anymore -- they are prefiltered by "all_facts"'
 
blanchet 
parents: 
51646 
diff
changeset
 | 
69  | 
val _ =  | 
| 
 
25acf689a53e
no need to filter tautologies anymore -- they are prefiltered by "all_facts"'
 
blanchet 
parents: 
51646 
diff
changeset
 | 
70  | 
      tracing ("Ran ATP: " ^
 | 
| 51649 | 71  | 
(case outcome of  | 
72  | 
NONE => "Success"  | 
|
73  | 
| SOME failure => string_for_failure failure))  | 
|
| 51646 | 74  | 
in outcome end  | 
| 
43566
 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 
blanchet 
parents: 
43499 
diff
changeset
 | 
75  | 
|
| 51646 | 76  | 
fun is_problem_line_reprovable ctxt format prelude axioms deps  | 
77  | 
(Formula (ident', _, phi, _, _)) =  | 
|
78  | 
is_none (run_atp ctxt format  | 
|
79  | 
((factsN, Formula (ident', Conjecture, phi, NONE, []) ::  | 
|
80  | 
map_filter (Symtab.lookup axioms) deps) ::  | 
|
81  | 
prelude))  | 
|
82  | 
| is_problem_line_reprovable _ _ _ _ _ _ = false  | 
|
83  | 
||
84  | 
fun inference_term _ [] = NONE  | 
|
85  | 
| inference_term check_infs ss =  | 
|
86  | 
    ATerm (("inference", []),
 | 
|
87  | 
        [ATerm (("checked_isabelle" |> not check_infs ? prefix "un", []), []),
 | 
|
88  | 
ATerm ((tptp_empty_list, []), []),  | 
|
89  | 
ATerm ((tptp_empty_list, []),  | 
|
90  | 
map (fn s => ATerm ((s, []), [])) ss)])  | 
|
91  | 
|> SOME  | 
|
92  | 
||
93  | 
fun add_inferences_to_problem_line ctxt format check_infs prelude axioms infers  | 
|
94  | 
(line as Formula ((ident, alt), Axiom, phi, NONE, tms)) =  | 
|
95  | 
let  | 
|
96  | 
val deps =  | 
|
97  | 
case these (AList.lookup (op =) infers ident) of  | 
|
98  | 
[] => []  | 
|
99  | 
| deps =>  | 
|
100  | 
if check_infs andalso  | 
|
101  | 
not (is_problem_line_reprovable ctxt format prelude axioms deps  | 
|
102  | 
line) then  | 
|
103  | 
[]  | 
|
104  | 
else  | 
|
105  | 
deps  | 
|
106  | 
in  | 
|
107  | 
Formula ((ident, alt), Lemma, phi, inference_term check_infs deps, tms)  | 
|
108  | 
end  | 
|
109  | 
| add_inferences_to_problem_line _ _ _ _ _ _ line = line  | 
|
110  | 
||
111  | 
fun add_inferences_to_problem ctxt format check_infs prelude infers problem =  | 
|
112  | 
let  | 
|
113  | 
fun add_if_axiom (axiom as Formula ((ident, _), Axiom, _, _, _)) =  | 
|
114  | 
Symtab.default (ident, axiom)  | 
|
115  | 
| add_if_axiom _ = I  | 
|
116  | 
val add_axioms_of_problem = fold (fold add_if_axiom o snd)  | 
|
117  | 
val axioms = Symtab.empty |> check_infs ? add_axioms_of_problem problem  | 
|
118  | 
in  | 
|
119  | 
map (apsnd (map (add_inferences_to_problem_line ctxt format check_infs  | 
|
120  | 
prelude axioms infers))) problem  | 
|
121  | 
end  | 
|
122  | 
||
123  | 
fun ident_of_problem_line (Class_Decl (ident, _, _)) = ident  | 
|
124  | 
| ident_of_problem_line (Type_Decl (ident, _, _)) = ident  | 
|
125  | 
| ident_of_problem_line (Sym_Decl (ident, _, _)) = ident  | 
|
126  | 
| ident_of_problem_line (Class_Memb (ident, _, _, _)) = ident  | 
|
127  | 
| ident_of_problem_line (Formula ((ident, _), _, _, _, _)) = ident  | 
|
128  | 
||
| 43499 | 129  | 
fun order_facts ord = sort (ord o pairself ident_of_problem_line)  | 
| 51646 | 130  | 
|
| 43499 | 131  | 
fun order_problem_facts _ [] = []  | 
132  | 
| order_problem_facts ord ((heading, lines) :: problem) =  | 
|
133  | 
if heading = factsN then (heading, order_facts ord lines) :: problem  | 
|
134  | 
else (heading, lines) :: order_problem_facts ord problem  | 
|
135  | 
||
| 45305 | 136  | 
(* A fairly random selection of types used for monomorphizing. *)  | 
137  | 
val ground_types =  | 
|
138  | 
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}, @{typ bool},
 | 
|
139  | 
   @{typ unit}]
 | 
|
140  | 
||
141  | 
fun ground_type_for_tvar _ [] tvar =  | 
|
142  | 
    raise TYPE ("ground_type_for_sorts", [TVar tvar], [])
 | 
|
143  | 
| ground_type_for_tvar thy (T :: Ts) tvar =  | 
|
144  | 
if can (Sign.typ_match thy (TVar tvar, T)) Vartab.empty then T  | 
|
145  | 
else ground_type_for_tvar thy Ts tvar  | 
|
146  | 
||
147  | 
fun monomorphize_term ctxt t =  | 
|
148  | 
let val thy = Proof_Context.theory_of ctxt in  | 
|
149  | 
t |> map_types (map_type_tvar (ground_type_for_tvar thy ground_types))  | 
|
150  | 
    handle TYPE _ => @{prop True}
 | 
|
151  | 
end  | 
|
152  | 
||
| 51633 | 153  | 
fun heading_sort_key heading =  | 
154  | 
if String.isPrefix "Relevant fact" heading then "_" ^ heading else heading  | 
|
155  | 
||
| 51646 | 156  | 
fun problem_for_theory ctxt thy format infer_policy type_enc =  | 
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
157  | 
let  | 
| 48299 | 158  | 
val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt  | 
| 
48315
 
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
 
blanchet 
parents: 
48301 
diff
changeset
 | 
159  | 
val type_enc =  | 
| 51632 | 160  | 
type_enc |> type_enc_from_string Non_Strict  | 
| 
48315
 
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
 
blanchet 
parents: 
48301 
diff
changeset
 | 
161  | 
|> adjust_type_enc format  | 
| 48131 | 162  | 
val mono = not (is_type_enc_polymorphic type_enc)  | 
| 
48438
 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 
blanchet 
parents: 
48406 
diff
changeset
 | 
163  | 
val facts =  | 
| 
50442
 
4f6a4d32522c
don't blacklist "case" theorems -- this causes problems in MaSh later
 
blanchet 
parents: 
48716 
diff
changeset
 | 
164  | 
Sledgehammer_Fact.all_facts (Proof_Context.init_global thy) true false  | 
| 
48530
 
d443166f9520
repaired accessibility chains generated by MaSh exporter + tuned one function out
 
blanchet 
parents: 
48438 
diff
changeset
 | 
165  | 
Symtab.empty [] [] css_table  | 
| 51633 | 166  | 
|> sort (Sledgehammer_MaSh.crude_thm_ord o pairself snd)  | 
| 51646 | 167  | 
val problem =  | 
| 
43576
 
ebeda6275027
don't export any metastrange or other nonatomizable formulas, since these don't help proving normal things, they are somewhat broken in the ATP output, and they are atypical
 
blanchet 
parents: 
43572 
diff
changeset
 | 
168  | 
facts  | 
| 45305 | 169  | 
|> map (fn ((_, loc), th) =>  | 
170  | 
((Thm.get_name_hint th, loc),  | 
|
171  | 
th |> prop_of |> mono ? monomorphize_term ctxt))  | 
|
| 
47946
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47912 
diff
changeset
 | 
172  | 
|> prepare_atp_problem ctxt format Axiom type_enc Exporter combsN false  | 
| 
 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 
blanchet 
parents: 
47912 
diff
changeset
 | 
173  | 
                             false true [] @{prop False}
 | 
| 
51647
 
25acf689a53e
no need to filter tautologies anymore -- they are prefiltered by "all_facts"'
 
blanchet 
parents: 
51646 
diff
changeset
 | 
174  | 
|> #1 |> sort_wrt (heading_sort_key o fst)  | 
| 51646 | 175  | 
val prelude = fst (split_last problem)  | 
| 
50735
 
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
 
blanchet 
parents: 
50521 
diff
changeset
 | 
176  | 
val name_tabs = Sledgehammer_Fact.build_name_tables Thm.get_name_hint facts  | 
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
177  | 
val infers =  | 
| 51646 | 178  | 
if infer_policy = No_Inferences then  | 
179  | 
[]  | 
|
180  | 
else  | 
|
181  | 
facts  | 
|
182  | 
|> map (fn (_, th) =>  | 
|
183  | 
(fact_name_of (Thm.get_name_hint th),  | 
|
184  | 
th |> Sledgehammer_Util.thms_in_proof (SOME name_tabs)  | 
|
185  | 
|> map fact_name_of))  | 
|
| 51650 | 186  | 
val all_problem_names =  | 
187  | 
problem |> maps (map ident_of_problem_line o snd)  | 
|
188  | 
|> distinct (op =)  | 
|
189  | 
val all_problem_name_set = Symtab.make (map (rpair ()) all_problem_names)  | 
|
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
190  | 
val infers =  | 
| 51650 | 191  | 
infers |> filter (Symtab.defined all_problem_name_set o fst)  | 
192  | 
|> map (apsnd (filter (Symtab.defined all_problem_name_set)))  | 
|
193  | 
val maybe_add_edge = perhaps o try o String_Graph.add_edge_acyclic  | 
|
| 43499 | 194  | 
val ordered_names =  | 
195  | 
String_Graph.empty  | 
|
| 51646 | 196  | 
|> fold (String_Graph.new_node o rpair ()) all_problem_names  | 
| 43499 | 197  | 
|> fold (fn (to, froms) =>  | 
| 51650 | 198  | 
fold (fn from => maybe_add_edge (from, to)) froms)  | 
| 
43566
 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 
blanchet 
parents: 
43499 
diff
changeset
 | 
199  | 
infers  | 
| 51650 | 200  | 
|> fold (fn (to, from) => maybe_add_edge (from, to))  | 
| 51646 | 201  | 
(tl all_problem_names ~~ fst (split_last all_problem_names))  | 
| 43499 | 202  | 
|> String_Graph.topological_order  | 
203  | 
val order_tab =  | 
|
204  | 
Symtab.empty  | 
|
205  | 
|> fold (Symtab.insert (op =))  | 
|
206  | 
(ordered_names ~~ (1 upto length ordered_names))  | 
|
207  | 
val name_ord = int_ord o pairself (the o Symtab.lookup order_tab)  | 
|
| 51646 | 208  | 
in  | 
209  | 
problem  | 
|
210  | 
|> (case format of  | 
|
211  | 
DFG _ => I  | 
|
212  | 
| _ => add_inferences_to_problem ctxt format  | 
|
213  | 
(infer_policy = Checked_Inferences) prelude infers)  | 
|
214  | 
|> order_problem_facts name_ord  | 
|
215  | 
end  | 
|
216  | 
||
217  | 
fun lines_for_problem ctxt format =  | 
|
218  | 
lines_for_atp_problem format (effective_term_order ctxt eN (* dummy *)) (K [])  | 
|
219  | 
||
220  | 
fun write_lines path ss =  | 
|
221  | 
let  | 
|
222  | 
val _ = File.write path ""  | 
|
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
223  | 
val _ = app (File.append path) ss  | 
| 
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
224  | 
in () end  | 
| 
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
225  | 
|
| 51646 | 226  | 
fun generate_atp_inference_file_for_theory ctxt thy format infer_policy type_enc  | 
227  | 
file_name =  | 
|
228  | 
let  | 
|
229  | 
val problem = problem_for_theory ctxt thy format infer_policy type_enc  | 
|
230  | 
val ss = lines_for_problem ctxt format problem  | 
|
231  | 
in write_lines (Path.explode file_name) ss end  | 
|
232  | 
||
233  | 
fun ap dir = Path.append dir o Path.explode  | 
|
234  | 
||
235  | 
fun chop_maximal_groups eq xs =  | 
|
236  | 
let  | 
|
237  | 
val rev_xs = rev xs  | 
|
238  | 
fun chop_group _ [] = []  | 
|
239  | 
| chop_group n (xs as x :: _) =  | 
|
240  | 
let  | 
|
241  | 
val n' = find_index (curry eq x) rev_xs  | 
|
242  | 
val (ws', xs') = chop (n - n') xs  | 
|
243  | 
in ws' :: chop_group n' xs' end  | 
|
244  | 
in chop_group (length xs) xs end  | 
|
245  | 
||
246  | 
fun theory_name_of_fact (Formula ((_, alt), _, _, _, _)) =  | 
|
247  | 
(case first_field Long_Name.separator alt of  | 
|
248  | 
NONE => alt  | 
|
249  | 
| SOME (thy, _) => thy)  | 
|
250  | 
| theory_name_of_fact _ = ""  | 
|
251  | 
||
252  | 
val problem_suffix = ".p"  | 
|
253  | 
val include_suffix = ".ax"  | 
|
254  | 
||
255  | 
val file_order_name = "FilesInProcessingOrder"  | 
|
256  | 
val problem_order_name = "ProblemsInProcessingOrder"  | 
|
257  | 
val problem_name = "problems"  | 
|
258  | 
val include_name = "incl"  | 
|
259  | 
val prelude_base_name = "prelude"  | 
|
260  | 
val prelude_name = prelude_base_name ^ include_suffix  | 
|
261  | 
||
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
262  | 
val encode_meta = Sledgehammer_MaSh.encode_str  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
263  | 
|
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
264  | 
val include_base_name_of_fact = encode_meta o theory_name_of_fact  | 
| 51646 | 265  | 
|
266  | 
fun include_line base_name =  | 
|
267  | 
  "include('" ^ include_name ^ "/" ^ base_name ^ include_suffix ^ "').\n"
 | 
|
268  | 
||
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
269  | 
val hol_base_name = encode_meta "HOL"  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
270  | 
|
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
271  | 
fun should_generate_problem thy base_name (Formula ((_, alt), _, _, _, _)) =  | 
| 51646 | 272  | 
case try (Global_Theory.get_thm thy) alt of  | 
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
273  | 
SOME th =>  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
274  | 
(* This is a crude hack to detect theorems stated and proved by the user (as  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
275  | 
opposed to those derived by various packages). In addition, we leave out  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
276  | 
everything in "HOL" as too basic to be interesting. *)  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
277  | 
Thm.legacy_get_kind th <> "" andalso base_name <> hol_base_name  | 
| 51646 | 278  | 
| NONE => false  | 
279  | 
||
280  | 
(* Convention: theoryname__goalname *)  | 
|
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
281  | 
fun problem_name_of base_name n alt =  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
282  | 
base_name ^ "__" ^ string_of_int n ^ "_" ^  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
283  | 
perhaps (try (unprefix (base_name ^ "_"))) alt ^ problem_suffix  | 
| 51646 | 284  | 
|
285  | 
fun generate_casc_lbt_isa_files_for_theory ctxt thy format infer_policy type_enc  | 
|
286  | 
dir_name =  | 
|
287  | 
let  | 
|
288  | 
val dir = Path.explode dir_name  | 
|
289  | 
val _ = Isabelle_System.mkdir dir  | 
|
290  | 
val file_order_path = ap dir file_order_name  | 
|
291  | 
val _ = File.write file_order_path ""  | 
|
292  | 
val problem_order_path = ap dir problem_order_name  | 
|
293  | 
val _ = File.write problem_order_path ""  | 
|
294  | 
val problem_dir = ap dir problem_name  | 
|
295  | 
val _ = Isabelle_System.mkdir problem_dir  | 
|
296  | 
val include_dir = ap problem_dir include_name  | 
|
297  | 
val _ = Isabelle_System.mkdir include_dir  | 
|
298  | 
val (prelude, groups) =  | 
|
299  | 
problem_for_theory ctxt thy format infer_policy type_enc  | 
|
300  | 
|> split_last  | 
|
301  | 
||> (snd  | 
|
302  | 
#> chop_maximal_groups (op = o pairself theory_name_of_fact)  | 
|
303  | 
#> map (`(include_base_name_of_fact o hd)))  | 
|
304  | 
fun write_prelude () =  | 
|
305  | 
let val ss = lines_for_problem ctxt format prelude in  | 
|
306  | 
File.append file_order_path (prelude_base_name ^ "\n");  | 
|
307  | 
write_lines (ap include_dir prelude_name) ss  | 
|
308  | 
end  | 
|
309  | 
fun write_include_file (base_name, facts) =  | 
|
310  | 
let  | 
|
311  | 
val name = base_name ^ include_suffix  | 
|
312  | 
val ss = lines_for_problem ctxt format [(factsN, facts)]  | 
|
313  | 
in  | 
|
314  | 
File.append file_order_path (base_name ^ "\n");  | 
|
315  | 
write_lines (ap include_dir name) ss  | 
|
316  | 
end  | 
|
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
317  | 
fun write_problem_files _ _ _ [] = ()  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
318  | 
| write_problem_files _ includes [] groups =  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
319  | 
write_problem_files 1 includes includes groups  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
320  | 
| write_problem_files n includes _ ((base_name, []) :: groups) =  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
321  | 
write_problem_files n (includes @ [include_line base_name]) [] groups  | 
| 
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
322  | 
| write_problem_files n includes seen  | 
| 51646 | 323  | 
((base_name, fact :: facts) :: groups) =  | 
324  | 
let val fact_s = tptp_string_for_line format fact in  | 
|
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
325  | 
(if should_generate_problem thy base_name fact then  | 
| 51646 | 326  | 
let  | 
327  | 
val (name, conj) =  | 
|
328  | 
(case fact of  | 
|
329  | 
Formula ((ident, alt), _, phi, _, _) =>  | 
|
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
330  | 
(problem_name_of base_name n (encode_meta alt),  | 
| 51646 | 331  | 
Formula ((ident, alt), Conjecture, phi, NONE, [])))  | 
332  | 
val conj_s = tptp_string_for_line format conj  | 
|
333  | 
in  | 
|
334  | 
File.append problem_order_path (name ^ "\n");  | 
|
335  | 
write_lines (ap problem_dir name) (seen @ [conj_s])  | 
|
336  | 
end  | 
|
337  | 
else  | 
|
338  | 
();  | 
|
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
339  | 
write_problem_files (n + 1) includes (seen @ [fact_s])  | 
| 51646 | 340  | 
((base_name, facts) :: groups))  | 
341  | 
end  | 
|
342  | 
val _ = write_prelude ()  | 
|
343  | 
val _ = app write_include_file groups  | 
|
| 
51652
 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 
blanchet 
parents: 
51650 
diff
changeset
 | 
344  | 
val _ = write_problem_files 1 [include_line prelude_base_name] [] groups  | 
| 51646 | 345  | 
in () end  | 
346  | 
||
| 
42602
 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 
blanchet 
parents:  
diff
changeset
 | 
347  | 
end;  |