| author | haftmann |
| Sat, 25 Jan 2014 23:50:49 +0100 | |
| changeset 55147 | bce3dbc11f95 |
| parent 51648 | 3e09226c3378 |
| child 58889 | 5b7a9633cfa8 |
| 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 |
|
|
51648
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
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:
42602
diff
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:
51646
diff
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}
|
|
51648
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
21 |
val thy = @{theory List}
|
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
22 |
val infer_policy = Unchecked_Inferences |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
23 |
*} |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
24 |
|
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
25 |
ML {*
|
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
26 |
if do_it then |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
27 |
"/tmp/axs_tc_native.dfg" |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
28 |
|> generate_atp_inference_file_for_theory ctxt thy (DFG Polymorphic) |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
29 |
infer_policy "tc_native" |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
30 |
else |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
31 |
() |
|
42602
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff
changeset
|
32 |
*} |
|
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff
changeset
|
33 |
|
|
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff
changeset
|
34 |
ML {*
|
|
51648
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
35 |
if do_it then |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
36 |
"/tmp/infs_poly_guards_query_query.tptp" |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
37 |
|> generate_atp_inference_file_for_theory ctxt thy FOF infer_policy |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
38 |
"poly_guards??" |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
39 |
else |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
40 |
() |
| 48301 | 41 |
*} |
42 |
||
|
51648
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
43 |
ML {*
|
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
44 |
if do_it then |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
45 |
"/tmp/infs_poly_tags_query_query.tptp" |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
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:
51646
diff
changeset
|
47 |
"poly_tags??" |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
48 |
else |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
49 |
() |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
50 |
*} |
| 51646 | 51 |
|
| 48301 | 52 |
ML {*
|
|
51648
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
53 |
if do_it then |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
54 |
"/tmp/casc_ltb_isa" |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
55 |
|> 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:
51646
diff
changeset
|
56 |
"poly_tags??" |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
57 |
else |
|
3e09226c3378
reverted accidental changes to theory file + updated wrt ML file
blanchet
parents:
51646
diff
changeset
|
58 |
() |
|
43400
3d03f4472883
use more appropriate type systems for ATP exporter
blanchet
parents:
43223
diff
changeset
|
59 |
*} |
|
3d03f4472883
use more appropriate type systems for ATP exporter
blanchet
parents:
43223
diff
changeset
|
60 |
|
|
42602
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff
changeset
|
61 |
end |