| author | paulson <lp15@cam.ac.uk> | 
| Tue, 31 Jan 2023 14:05:16 +0000 | |
| changeset 77140 | 9a60c1759543 | 
| parent 74902 | ece4f07ebb04 | 
| permissions | -rw-r--r-- | 
| 46321 | 1 | (* Title: HOL/TPTP/ATP_Theory_Export.thy | 
| 2 | Author: Jasmin Blanchette, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 63167 | 5 | section \<open>ATP Theory Exporter\<close> | 
| 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 | |
| 69605 | 11 | ML_file \<open>atp_theory_export.ML\<close> | 
| 48891 | 12 | |
| 63167 | 13 | ML \<open> | 
| 48375 | 14 | open ATP_Problem | 
| 15 | open ATP_Theory_Export | |
| 63167 | 16 | \<close> | 
| 45305 | 17 | |
| 63167 | 18 | ML \<open> | 
| 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 *) | 
| 69597 | 20 | val ctxt = \<^context> | 
| 21 | val thy = \<^theory>\<open>Complex_Main\<close> | |
| 61323 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 22 | val infer_policy = (* Unchecked_Inferences *) No_Inferences | 
| 63167 | 23 | \<close> | 
| 61323 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 24 | |
| 63167 | 25 | ML \<open> | 
| 61323 
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" | 
| 72591 | 28 | |> generate_casc_lbt_isa_files_for_theory ctxt thy | 
| 74902 | 29 |      (THF (Polymorphic, {with_ite = false, with_let = false}, THF_Without_Choice)) infer_policy
 | 
| 30 | "poly_native_higher" | |
| 61323 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 31 | else | 
| 
99b3a17a7eab
extended theory exporter to also export MePo-selected facts
 blanchet parents: 
58889diff
changeset | 32 | () | 
| 63167 | 33 | \<close> | 
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 34 | |
| 63167 | 35 | ML \<open> | 
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 36 | if do_it then | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 37 | "/tmp/axs_tc_native.dfg" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 38 | |> 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 | 39 | infer_policy "tc_native" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 40 | else | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 41 | () | 
| 63167 | 42 | \<close> | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 43 | |
| 63167 | 44 | ML \<open> | 
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 45 | if do_it then | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 46 | "/tmp/infs_poly_guards_query_query.tptp" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 47 | |> 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 | 48 | "poly_guards??" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 49 | else | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 50 | () | 
| 63167 | 51 | \<close> | 
| 48301 | 52 | |
| 63167 | 53 | ML \<open> | 
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 54 | if do_it then | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 55 | "/tmp/infs_poly_tags_query_query.tptp" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 56 | |> 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 | 57 | "poly_tags??" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 58 | else | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 59 | () | 
| 63167 | 60 | \<close> | 
| 51646 | 61 | |
| 63167 | 62 | ML \<open> | 
| 51648 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 63 | if do_it then | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 64 | "/tmp/casc_ltb_isa" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 65 | |> 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 | 66 | "poly_tags??" | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 67 | else | 
| 
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
 blanchet parents: 
51646diff
changeset | 68 | () | 
| 63167 | 69 | \<close> | 
| 43400 
3d03f4472883
use more appropriate type systems for ATP exporter
 blanchet parents: 
43223diff
changeset | 70 | |
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 71 | end |