| author | hoelzl |
| Thu, 17 Jan 2013 11:59:12 +0100 | |
| changeset 50936 | b28f258ebc1a |
| parent 50925 | dfc0177384f9 |
| child 50954 | 7bc58677860e |
| permissions | -rw-r--r-- |
| 48234 | 1 |
(* Title: HOL/TPTP/MaSh_Export.thy |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
3 |
*) |
|
4 |
||
5 |
header {* MaSh Exporter *}
|
|
6 |
||
7 |
theory MaSh_Export |
|
|
50350
136d5318b1fe
tuned MaSh exporter -- and don't make temp directories unless explicitly told so
blanchet
parents:
50349
diff
changeset
|
8 |
imports Main |
| 48234 | 9 |
begin |
10 |
||
| 48891 | 11 |
ML_file "mash_export.ML" |
12 |
||
| 48245 | 13 |
sledgehammer_params |
| 50519 | 14 |
[provers = spass, max_facts = 32, strict, dont_slice, type_enc = mono_native, |
| 50925 | 15 |
lam_trans = lifting, timeout = 2, dont_preplay, minimize] |
|
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
16 |
|
| 50925 | 17 |
declare [[sledgehammer_instantiate_inducts = false]] |
|
50434
960a3429615c
more MaSh tweaking -- in particular, export the same facts in "MaSh_Export" as are later tried in "MaSh_Eval"
blanchet
parents:
50411
diff
changeset
|
18 |
|
| 48234 | 19 |
ML {*
|
| 50519 | 20 |
!Multithreading.max_threads |
21 |
*} |
|
22 |
||
23 |
ML {*
|
|
| 48234 | 24 |
open MaSh_Export |
25 |
*} |
|
26 |
||
27 |
ML {*
|
|
| 48333 | 28 |
val do_it = false (* switch to "true" to generate the files *) |
|
50350
136d5318b1fe
tuned MaSh exporter -- and don't make temp directories unless explicitly told so
blanchet
parents:
50349
diff
changeset
|
29 |
val thys = [@{theory List}]
|
| 48333 | 30 |
val params as {provers, ...} = Sledgehammer_Isar.default_params @{context} []
|
31 |
val prover = hd provers |
|
| 50925 | 32 |
val max_suggestions = 1536 |
| 50513 | 33 |
val dir = "List" |
| 50349 | 34 |
val prefix = "/tmp/" ^ dir ^ "/" |
| 50337 | 35 |
*} |
36 |
||
37 |
ML {*
|
|
|
50350
136d5318b1fe
tuned MaSh exporter -- and don't make temp directories unless explicitly told so
blanchet
parents:
50349
diff
changeset
|
38 |
if do_it then |
|
136d5318b1fe
tuned MaSh exporter -- and don't make temp directories unless explicitly told so
blanchet
parents:
50349
diff
changeset
|
39 |
Isabelle_System.mkdir (Path.explode prefix) |
|
136d5318b1fe
tuned MaSh exporter -- and don't make temp directories unless explicitly told so
blanchet
parents:
50349
diff
changeset
|
40 |
else |
|
136d5318b1fe
tuned MaSh exporter -- and don't make temp directories unless explicitly told so
blanchet
parents:
50349
diff
changeset
|
41 |
() |
|
48246
fb11c09d7729
add Isabelle dependencies to tweak relevance filter
blanchet
parents:
48245
diff
changeset
|
42 |
*} |
| 48531 | 43 |
|
|
48246
fb11c09d7729
add Isabelle dependencies to tweak relevance filter
blanchet
parents:
48245
diff
changeset
|
44 |
ML {*
|
| 48245 | 45 |
if do_it then |
| 50349 | 46 |
generate_accessibility @{context} thys false (prefix ^ "mash_accessibility")
|
| 48245 | 47 |
else |
48 |
() |
|
49 |
*} |
|
50 |
||
51 |
ML {*
|
|
52 |
if do_it then |
|
| 50349 | 53 |
generate_features @{context} prover thys false (prefix ^ "mash_features")
|
| 48245 | 54 |
else |
55 |
() |
|
|
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
56 |
*} |
|
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
57 |
|
|
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48235
diff
changeset
|
58 |
ML {*
|
| 48245 | 59 |
if do_it then |
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50434
diff
changeset
|
60 |
generate_isar_dependencies @{context} thys false
|
|
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50434
diff
changeset
|
61 |
(prefix ^ "mash_dependencies") |
| 48333 | 62 |
else |
63 |
() |
|
64 |
*} |
|
65 |
||
66 |
ML {*
|
|
67 |
if do_it then |
|
| 50411 | 68 |
generate_isar_commands @{context} prover thys (prefix ^ "mash_commands")
|
| 48333 | 69 |
else |
70 |
() |
|
71 |
*} |
|
72 |
||
73 |
ML {*
|
|
| 50816 | 74 |
if do_it then |
| 50906 | 75 |
generate_mepo_suggestions @{context} params (1, NONE) thys max_suggestions
|
| 50814 | 76 |
(prefix ^ "mepo_suggestions") |
77 |
else |
|
78 |
() |
|
79 |
*} |
|
80 |
||
81 |
ML {*
|
|
| 48333 | 82 |
if do_it then |
| 50814 | 83 |
generate_mesh_suggestions max_suggestions (prefix ^ "mash_suggestions") |
84 |
(prefix ^ "mepo_suggestions") (prefix ^ "mesh_suggestions") |
|
| 48245 | 85 |
else |
86 |
() |
|
87 |
*} |
|
88 |
||
89 |
ML {*
|
|
90 |
if do_it then |
|
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50519
diff
changeset
|
91 |
generate_prover_dependencies @{context} params (1, NONE) thys false
|
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50434
diff
changeset
|
92 |
(prefix ^ "mash_prover_dependencies") |
| 48245 | 93 |
else |
94 |
() |
|
| 48234 | 95 |
*} |
96 |
||
| 50411 | 97 |
ML {*
|
98 |
if do_it then |
|
|
50559
89c0d2f13cca
MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents:
50519
diff
changeset
|
99 |
generate_prover_commands @{context} params (1, NONE) thys
|
|
50484
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents:
50434
diff
changeset
|
100 |
(prefix ^ "mash_prover_commands") |
| 50411 | 101 |
else |
102 |
() |
|
103 |
*} |
|
104 |
||
| 48234 | 105 |
end |