| author | wenzelm | 
| Fri, 08 Apr 2016 20:52:40 +0200 | |
| changeset 62914 | 930a30c1a9af | 
| parent 61323 | 99b3a17a7eab | 
| child 63167 | 0909deb8059b | 
| permissions | -rw-r--r-- | 
| 46321 | 1 | (* Title: HOL/TPTP/ATP_Theory_Export.thy | 
| 2 | Author: Jasmin Blanchette, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 58889 | 5 | section {* ATP Theory Exporter *}
 | 
| 46321 | 6 | |
| 7 | theory ATP_Theory_Export | |
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 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 {*
 | |
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 19 | val do_it = false (* switch to "true" to generate the files *) | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 20 | val ctxt = @{context}
 | 
| 61323 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 21 | val thy = @{theory Complex_Main}
 | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 22 | val infer_policy = (* Unchecked_Inferences *) No_Inferences | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 23 | *} | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 24 | |
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 25 | ML {*
 | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 26 | if do_it then | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 27 | "/tmp/isa_filter" | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 28 | |> generate_casc_lbt_isa_files_for_theory ctxt thy (THF (Polymorphic, THF_Without_Choice)) | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 29 | infer_policy "poly_native_higher" | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 30 | else | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 31 | () | 
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 32 | *} | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 33 | |
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 34 | ML {*
 | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 35 | if do_it then | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 36 | "/tmp/axs_tc_native.dfg" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 37 | |> generate_atp_inference_file_for_theory ctxt thy (DFG Polymorphic) | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 38 | infer_policy "tc_native" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 39 | else | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 40 | () | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 41 | *} | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 42 | |
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 43 | ML {*
 | 
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 44 | if do_it then | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 45 | "/tmp/infs_poly_guards_query_query.tptp" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 46 | |> generate_atp_inference_file_for_theory ctxt thy FOF infer_policy | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 47 | "poly_guards??" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 48 | else | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 49 | () | 
| 48301 | 50 | *} | 
| 51 | ||
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 52 | ML {*
 | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 53 | if do_it then | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 54 | "/tmp/infs_poly_tags_query_query.tptp" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 55 | |> generate_atp_inference_file_for_theory ctxt thy FOF infer_policy | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 56 | "poly_tags??" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 57 | else | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 58 | () | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 59 | *} | 
| 51646 | 60 | |
| 48301 | 61 | ML {*
 | 
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 62 | if do_it then | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 63 | "/tmp/casc_ltb_isa" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 64 | |> generate_casc_lbt_isa_files_for_theory ctxt thy FOF infer_policy | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 65 | "poly_tags??" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 66 | else | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 67 | () | 
| 43400 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 68 | *} | 
| 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 69 | |
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 70 | end |