| author | blanchet | 
| Fri, 21 Dec 2012 14:35:29 +0100 | |
| changeset 50610 | d9c4fbbb0c11 | 
| parent 48891 | c0eafbd55de3 | 
| child 51645 | 86e8c87e1f1b | 
| permissions | -rw-r--r-- | 
| 46321 | 1 | (* Title: HOL/TPTP/ATP_Theory_Export.thy | 
| 2 | Author: Jasmin Blanchette, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* ATP Theory Exporter *}
 | |
| 6 | ||
| 7 | theory ATP_Theory_Export | |
| 48302 | 8 | imports Complex_Main | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 9 | begin | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 10 | |
| 48891 | 11 | ML_file "atp_theory_export.ML" | 
| 12 | ||
| 42607 
c8673078f915
fix ROOT.ML and handle "readable_names" reference slightly more cleanly
 blanchet parents: 
42602diff
changeset | 13 | ML {*
 | 
| 48375 | 14 | open ATP_Problem | 
| 15 | open ATP_Theory_Export | |
| 45305 | 16 | *} | 
| 17 | ||
| 18 | ML {*
 | |
| 48375 | 19 | val do_it = false (* switch to "true" to generate the files *) | 
| 20 | val thy = @{theory List}
 | |
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 21 | val ctxt = @{context}
 | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 22 | *} | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 23 | |
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 24 | ML {*
 | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42607diff
changeset | 25 | if do_it then | 
| 48301 | 26 | "/tmp/axs_tc_native.dfg" | 
| 27 | |> generate_atp_inference_file_for_theory ctxt thy (DFG Polymorphic) | |
| 28 | "tc_native" | |
| 29 | else | |
| 30 | () | |
| 31 | *} | |
| 32 | ||
| 33 | ML {*
 | |
| 34 | if do_it then | |
| 48213 
d20add034f64
first go at generating files for MaSh (machine-learning Sledgehammer)
 blanchet parents: 
48145diff
changeset | 35 | "/tmp/infs_poly_guards_query_query.tptp" | 
| 48234 | 36 | |> generate_atp_inference_file_for_theory ctxt thy FOF | 
| 48217 
8994afe09c18
more precise dependencies -- eliminate tautologies
 blanchet parents: 
48216diff
changeset | 37 | "poly_guards_query_query" | 
| 43400 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 38 | else | 
| 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 39 | () | 
| 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 40 | *} | 
| 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 41 | |
| 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 42 | ML {*
 | 
| 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 43 | if do_it then | 
| 48145 | 44 | "/tmp/infs_poly_tags_query_query.tptp" | 
| 48234 | 45 | |> generate_atp_inference_file_for_theory ctxt thy FOF | 
| 48217 
8994afe09c18
more precise dependencies -- eliminate tautologies
 blanchet parents: 
48216diff
changeset | 46 | "poly_tags_query_query" | 
| 45305 | 47 | else | 
| 48 | () | |
| 49 | *} | |
| 50 | ||
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 51 | end |