| author | wenzelm | 
| Mon, 13 Oct 2014 21:41:29 +0200 | |
| changeset 58665 | 50b229a5a097 | 
| parent 57676 | d53b1f876afb | 
| child 58922 | 1f500b18c4c6 | 
| permissions | -rw-r--r-- | 
| 46321 | 1 | (* Title: HOL/TPTP/atp_theory_export.ML | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 2 | Author: Jasmin Blanchette, TU Muenchen | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 3 | Copyright 2011 | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 4 | |
| 48234 | 5 | Export Isabelle theories as first-order TPTP inferences. | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 6 | *) | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 7 | |
| 46321 | 8 | signature ATP_THEORY_EXPORT = | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 9 | sig | 
| 45305 | 10 | type atp_format = ATP_Problem.atp_format | 
| 11 | ||
| 51646 | 12 | datatype inference_policy = | 
| 13 | No_Inferences | Unchecked_Inferences | Checked_Inferences | |
| 14 | ||
| 48234 | 15 | val generate_atp_inference_file_for_theory : | 
| 51646 | 16 | Proof.context -> theory -> atp_format -> inference_policy -> string | 
| 17 | -> string -> unit | |
| 18 | val generate_casc_lbt_isa_files_for_theory : | |
| 19 | Proof.context -> theory -> atp_format -> inference_policy -> string | |
| 20 | -> string -> unit | |
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 21 | end; | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 22 | |
| 48234 | 23 | structure ATP_Theory_Export : ATP_THEORY_EXPORT = | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 24 | struct | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 25 | |
| 43479 
5af1abc13c1f
only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
 blanchet parents: 
43468diff
changeset | 26 | open ATP_Problem | 
| 43566 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43499diff
changeset | 27 | open ATP_Proof | 
| 46320 | 28 | open ATP_Problem_Generate | 
| 43566 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43499diff
changeset | 29 | open ATP_Systems | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 30 | |
| 51646 | 31 | datatype inference_policy = | 
| 32 | No_Inferences | Unchecked_Inferences | Checked_Inferences | |
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 33 | |
| 53980 
7e6a82c593f4
made SML/NJ happy (NB: toplevel ML environment is unmanaged);
 wenzelm parents: 
53586diff
changeset | 34 | val prefix = Library.prefix | 
| 51646 | 35 | val fact_name_of = prefix fact_prefix o ascii_of | 
| 43479 
5af1abc13c1f
only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
 blanchet parents: 
43468diff
changeset | 36 | |
| 54197 
994ebb795b75
use definitions for LEO-II as well -- this simplifies the code and matches some users' expectations
 blanchet parents: 
53980diff
changeset | 37 | fun atp_of_format (THF (Polymorphic, _)) = dummy_thfN | 
| 
994ebb795b75
use definitions for LEO-II as well -- this simplifies the code and matches some users' expectations
 blanchet parents: 
53980diff
changeset | 38 | | atp_of_format (THF (Monomorphic, _)) = satallaxN | 
| 54788 
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
 blanchet parents: 
54197diff
changeset | 39 | | atp_of_format (DFG Polymorphic) = spass_pirateN | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 40 | | atp_of_format (DFG Monomorphic) = spassN | 
| 52995 
ab98feb66684
Vampire 3.0 requires types to be declared -- make it happy (and get rid of "implicit" types since only Satallax seems to support them anymore)
 blanchet parents: 
52754diff
changeset | 41 | | atp_of_format (TFF Polymorphic) = alt_ergoN | 
| 
ab98feb66684
Vampire 3.0 requires types to be declared -- make it happy (and get rid of "implicit" types since only Satallax seems to support them anymore)
 blanchet parents: 
52754diff
changeset | 42 | | atp_of_format (TFF Monomorphic) = vampireN | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 43 | | atp_of_format FOF = eN | 
| 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 44 | | atp_of_format CNF_UEQ = waldmeisterN | 
| 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 45 | | atp_of_format CNF = eN | 
| 48318 | 46 | |
| 51646 | 47 | val atp_timeout = seconds 0.5 | 
| 48 | ||
| 49 | fun run_atp ctxt format problem = | |
| 43566 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43499diff
changeset | 50 | let | 
| 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43499diff
changeset | 51 | val thy = Proof_Context.theory_of ctxt | 
| 47055 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47038diff
changeset | 52 | val prob_file = File.tmp_path (Path.explode "prob") | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 53 | val atp = atp_of_format format | 
| 57676 | 54 |     val {exec, arguments, proof_delims, known_failures, ...} = get_atp thy atp ()
 | 
| 47038 
2409b484e1cc
continued implementation of term ordering attributes
 blanchet parents: 
46734diff
changeset | 55 | val ord = effective_term_order ctxt atp | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 56 | val _ = problem |> lines_of_atp_problem format ord (K []) | 
| 46442 
1e07620d724c
added possibility of generating KBO weights to DFG problems
 blanchet parents: 
46409diff
changeset | 57 | |> File.write_list prob_file | 
| 57676 | 58 | val exec = exec false | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48324diff
changeset | 59 | val path = getenv (List.last (fst exec)) ^ "/" ^ List.last (snd exec) | 
| 43566 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43499diff
changeset | 60 | val command = | 
| 50927 | 61 | File.shell_path (Path.explode path) ^ " " ^ | 
| 51646 | 62 | arguments ctxt false "" atp_timeout (File.shell_path prob_file) | 
| 50927 | 63 | (ord, K [], K []) | 
| 51646 | 64 | val outcome = | 
| 65 | TimeLimit.timeLimit atp_timeout Isabelle_System.bash_output command | |
| 66 | |> fst | |
| 67 | |> extract_tstplike_proof_and_outcome false proof_delims known_failures | |
| 68 | |> snd | |
| 69 | handle TimeLimit.TimeOut => SOME TimedOut | |
| 51647 
25acf689a53e
no need to filter tautologies anymore -- they are prefiltered by "all_facts"'
 blanchet parents: 
51646diff
changeset | 70 | val _ = | 
| 
25acf689a53e
no need to filter tautologies anymore -- they are prefiltered by "all_facts"'
 blanchet parents: 
51646diff
changeset | 71 |       tracing ("Ran ATP: " ^
 | 
| 51649 | 72 | (case outcome of | 
| 73 | NONE => "Success" | |
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
52995diff
changeset | 74 | | SOME failure => string_of_atp_failure failure)) | 
| 51646 | 75 | in outcome end | 
| 43566 
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
 blanchet parents: 
43499diff
changeset | 76 | |
| 51646 | 77 | fun is_problem_line_reprovable ctxt format prelude axioms deps | 
| 78 | (Formula (ident', _, phi, _, _)) = | |
| 79 | is_none (run_atp ctxt format | |
| 80 | ((factsN, Formula (ident', Conjecture, phi, NONE, []) :: | |
| 81 | map_filter (Symtab.lookup axioms) deps) :: | |
| 82 | prelude)) | |
| 83 | | is_problem_line_reprovable _ _ _ _ _ _ = false | |
| 84 | ||
| 85 | fun inference_term _ [] = NONE | |
| 86 | | inference_term check_infs ss = | |
| 87 |     ATerm (("inference", []),
 | |
| 88 |         [ATerm (("checked_isabelle" |> not check_infs ? prefix "un", []), []),
 | |
| 89 | ATerm ((tptp_empty_list, []), []), | |
| 90 | ATerm ((tptp_empty_list, []), | |
| 91 | map (fn s => ATerm ((s, []), [])) ss)]) | |
| 92 | |> SOME | |
| 93 | ||
| 94 | fun add_inferences_to_problem_line ctxt format check_infs prelude axioms infers | |
| 95 | (line as Formula ((ident, alt), Axiom, phi, NONE, tms)) = | |
| 96 | let | |
| 97 | val deps = | |
| 98 | case these (AList.lookup (op =) infers ident) of | |
| 99 | [] => [] | |
| 100 | | deps => | |
| 101 | if check_infs andalso | |
| 102 | not (is_problem_line_reprovable ctxt format prelude axioms deps | |
| 103 | line) then | |
| 104 | [] | |
| 105 | else | |
| 106 | deps | |
| 107 | in | |
| 108 | Formula ((ident, alt), Lemma, phi, inference_term check_infs deps, tms) | |
| 109 | end | |
| 110 | | add_inferences_to_problem_line _ _ _ _ _ _ line = line | |
| 111 | ||
| 112 | fun add_inferences_to_problem ctxt format check_infs prelude infers problem = | |
| 113 | let | |
| 114 | fun add_if_axiom (axiom as Formula ((ident, _), Axiom, _, _, _)) = | |
| 115 | Symtab.default (ident, axiom) | |
| 116 | | add_if_axiom _ = I | |
| 117 | val add_axioms_of_problem = fold (fold add_if_axiom o snd) | |
| 118 | val axioms = Symtab.empty |> check_infs ? add_axioms_of_problem problem | |
| 119 | in | |
| 120 | map (apsnd (map (add_inferences_to_problem_line ctxt format check_infs | |
| 121 | prelude axioms infers))) problem | |
| 122 | end | |
| 123 | ||
| 124 | fun ident_of_problem_line (Class_Decl (ident, _, _)) = ident | |
| 125 | | ident_of_problem_line (Type_Decl (ident, _, _)) = ident | |
| 126 | | ident_of_problem_line (Sym_Decl (ident, _, _)) = ident | |
| 127 | | ident_of_problem_line (Class_Memb (ident, _, _, _)) = ident | |
| 128 | | ident_of_problem_line (Formula ((ident, _), _, _, _, _)) = ident | |
| 129 | ||
| 43499 | 130 | fun order_facts ord = sort (ord o pairself ident_of_problem_line) | 
| 51646 | 131 | |
| 43499 | 132 | fun order_problem_facts _ [] = [] | 
| 133 | | order_problem_facts ord ((heading, lines) :: problem) = | |
| 134 | if heading = factsN then (heading, order_facts ord lines) :: problem | |
| 135 | else (heading, lines) :: order_problem_facts ord problem | |
| 136 | ||
| 45305 | 137 | (* A fairly random selection of types used for monomorphizing. *) | 
| 138 | val ground_types = | |
| 139 |   [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}, @{typ bool},
 | |
| 140 |    @{typ unit}]
 | |
| 141 | ||
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 142 | fun ground_type_of_tvar _ [] tvar = | 
| 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 143 |     raise TYPE ("ground_type_of_tvar", [TVar tvar], [])
 | 
| 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 144 | | ground_type_of_tvar thy (T :: Ts) tvar = | 
| 45305 | 145 | if can (Sign.typ_match thy (TVar tvar, T)) Vartab.empty then T | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 146 | else ground_type_of_tvar thy Ts tvar | 
| 45305 | 147 | |
| 148 | fun monomorphize_term ctxt t = | |
| 149 | let val thy = Proof_Context.theory_of ctxt in | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 150 | t |> map_types (map_type_tvar (ground_type_of_tvar thy ground_types)) | 
| 45305 | 151 |     handle TYPE _ => @{prop True}
 | 
| 152 | end | |
| 153 | ||
| 51633 | 154 | fun heading_sort_key heading = | 
| 155 | if String.isPrefix "Relevant fact" heading then "_" ^ heading else heading | |
| 156 | ||
| 57306 
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
 blanchet parents: 
57268diff
changeset | 157 | val max_dependencies = 100 | 
| 
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
 blanchet parents: 
57268diff
changeset | 158 | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 159 | fun problem_of_theory ctxt thy format infer_policy type_enc = | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 160 | let | 
| 48299 | 161 | val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt | 
| 48315 
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
 blanchet parents: 
48301diff
changeset | 162 | val type_enc = | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51998diff
changeset | 163 | type_enc |> type_enc_of_string Non_Strict | 
| 48315 
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
 blanchet parents: 
48301diff
changeset | 164 | |> adjust_type_enc format | 
| 48131 | 165 | val mono = not (is_type_enc_polymorphic type_enc) | 
| 48438 
3e45c98fe127
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
 blanchet parents: 
48406diff
changeset | 166 | val facts = | 
| 50442 
4f6a4d32522c
don't blacklist "case" theorems -- this causes problems in MaSh later
 blanchet parents: 
48716diff
changeset | 167 | Sledgehammer_Fact.all_facts (Proof_Context.init_global thy) true false | 
| 48530 
d443166f9520
repaired accessibility chains generated by MaSh exporter + tuned one function out
 blanchet parents: 
48438diff
changeset | 168 | Symtab.empty [] [] css_table | 
| 51633 | 169 | |> sort (Sledgehammer_MaSh.crude_thm_ord o pairself snd) | 
| 51646 | 170 | val problem = | 
| 43576 
ebeda6275027
don't export any metastrange or other nonatomizable formulas, since these don't help proving normal things, they are somewhat broken in the ATP output, and they are atypical
 blanchet parents: 
43572diff
changeset | 171 | facts | 
| 45305 | 172 | |> map (fn ((_, loc), th) => | 
| 57268 | 173 | ((Thm.get_name_hint th, loc), th |> prop_of |> mono ? monomorphize_term ctxt)) | 
| 174 | |> generate_atp_problem ctxt format Axiom type_enc Exporter combsN false false true [] | |
| 175 |         @{prop False}
 | |
| 51647 
25acf689a53e
no need to filter tautologies anymore -- they are prefiltered by "all_facts"'
 blanchet parents: 
51646diff
changeset | 176 | |> #1 |> sort_wrt (heading_sort_key o fst) | 
| 51646 | 177 | val prelude = fst (split_last problem) | 
| 50735 
6b232d76cbc9
refined class handling, to prevent cycles in fact graph
 blanchet parents: 
50521diff
changeset | 178 | val name_tabs = Sledgehammer_Fact.build_name_tables Thm.get_name_hint facts | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 179 | val infers = | 
| 51646 | 180 | if infer_policy = No_Inferences then | 
| 181 | [] | |
| 182 | else | |
| 183 | facts | |
| 184 | |> map (fn (_, th) => | |
| 185 | (fact_name_of (Thm.get_name_hint th), | |
| 57306 
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
 blanchet parents: 
57268diff
changeset | 186 | th |> Sledgehammer_Util.thms_in_proof max_dependencies (SOME name_tabs) | 
| 
ff10067b2248
optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
 blanchet parents: 
57268diff
changeset | 187 | |> these |> map fact_name_of)) | 
| 51650 | 188 | val all_problem_names = | 
| 57307 | 189 | problem |> maps (map ident_of_problem_line o snd) |> distinct (op =) | 
| 51650 | 190 | val all_problem_name_set = Symtab.make (map (rpair ()) all_problem_names) | 
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 191 | val infers = | 
| 51650 | 192 | infers |> filter (Symtab.defined all_problem_name_set o fst) | 
| 193 | |> map (apsnd (filter (Symtab.defined all_problem_name_set))) | |
| 194 | val maybe_add_edge = perhaps o try o String_Graph.add_edge_acyclic | |
| 43499 | 195 | val ordered_names = | 
| 196 | String_Graph.empty | |
| 51646 | 197 | |> fold (String_Graph.new_node o rpair ()) all_problem_names | 
| 57307 | 198 | |> fold (fn (to, froms) => fold (fn from => maybe_add_edge (from, to)) froms) infers | 
| 51650 | 199 | |> fold (fn (to, from) => maybe_add_edge (from, to)) | 
| 57307 | 200 | (tl all_problem_names ~~ fst (split_last all_problem_names)) | 
| 43499 | 201 | |> String_Graph.topological_order | 
| 202 | val order_tab = | |
| 203 | Symtab.empty | |
| 57307 | 204 | |> fold (Symtab.insert (op =)) (ordered_names ~~ (1 upto length ordered_names)) | 
| 43499 | 205 | val name_ord = int_ord o pairself (the o Symtab.lookup order_tab) | 
| 51646 | 206 | in | 
| 207 | problem | |
| 208 | |> (case format of | |
| 209 | DFG _ => I | |
| 210 | | _ => add_inferences_to_problem ctxt format | |
| 211 | (infer_policy = Checked_Inferences) prelude infers) | |
| 212 | |> order_problem_facts name_ord | |
| 213 | end | |
| 214 | ||
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 215 | fun lines_of_problem ctxt format = | 
| 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 216 | lines_of_atp_problem format (effective_term_order ctxt eN (* dummy *)) (K []) | 
| 51646 | 217 | |
| 218 | fun write_lines path ss = | |
| 219 | let | |
| 220 | val _ = File.write path "" | |
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 221 | val _ = app (File.append path) ss | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 222 | in () end | 
| 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 223 | |
| 51646 | 224 | fun generate_atp_inference_file_for_theory ctxt thy format infer_policy type_enc | 
| 225 | file_name = | |
| 226 | let | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 227 | val problem = problem_of_theory ctxt thy format infer_policy type_enc | 
| 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 228 | val ss = lines_of_problem ctxt format problem | 
| 51646 | 229 | in write_lines (Path.explode file_name) ss end | 
| 230 | ||
| 231 | fun ap dir = Path.append dir o Path.explode | |
| 232 | ||
| 233 | fun chop_maximal_groups eq xs = | |
| 234 | let | |
| 235 | val rev_xs = rev xs | |
| 236 | fun chop_group _ [] = [] | |
| 237 | | chop_group n (xs as x :: _) = | |
| 238 | let | |
| 239 | val n' = find_index (curry eq x) rev_xs | |
| 240 | val (ws', xs') = chop (n - n') xs | |
| 241 | in ws' :: chop_group n' xs' end | |
| 242 | in chop_group (length xs) xs end | |
| 243 | ||
| 244 | fun theory_name_of_fact (Formula ((_, alt), _, _, _, _)) = | |
| 245 | (case first_field Long_Name.separator alt of | |
| 246 | NONE => alt | |
| 247 | | SOME (thy, _) => thy) | |
| 248 | | theory_name_of_fact _ = "" | |
| 249 | ||
| 250 | val problem_suffix = ".p" | |
| 251 | val include_suffix = ".ax" | |
| 252 | ||
| 253 | val file_order_name = "FilesInProcessingOrder" | |
| 254 | val problem_order_name = "ProblemsInProcessingOrder" | |
| 255 | val problem_name = "problems" | |
| 256 | val include_name = "incl" | |
| 257 | val prelude_base_name = "prelude" | |
| 258 | val prelude_name = prelude_base_name ^ include_suffix | |
| 259 | ||
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 260 | val encode_meta = Sledgehammer_MaSh.encode_str | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 261 | |
| 53980 
7e6a82c593f4
made SML/NJ happy (NB: toplevel ML environment is unmanaged);
 wenzelm parents: 
53586diff
changeset | 262 | fun include_base_name_of_fact x = encode_meta (theory_name_of_fact x) | 
| 51646 | 263 | |
| 264 | fun include_line base_name = | |
| 265 |   "include('" ^ include_name ^ "/" ^ base_name ^ include_suffix ^ "').\n"
 | |
| 266 | ||
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 267 | val hol_base_name = encode_meta "HOL" | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 268 | |
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 269 | fun should_generate_problem thy base_name (Formula ((_, alt), _, _, _, _)) = | 
| 51646 | 270 | case try (Global_Theory.get_thm thy) alt of | 
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 271 | SOME th => | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 272 | (* This is a crude hack to detect theorems stated and proved by the user (as | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 273 | opposed to those derived by various packages). In addition, we leave out | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 274 | everything in "HOL" as too basic to be interesting. *) | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 275 | Thm.legacy_get_kind th <> "" andalso base_name <> hol_base_name | 
| 51646 | 276 | | NONE => false | 
| 277 | ||
| 278 | (* Convention: theoryname__goalname *) | |
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 279 | fun problem_name_of base_name n alt = | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 280 | base_name ^ "__" ^ string_of_int n ^ "_" ^ | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 281 | perhaps (try (unprefix (base_name ^ "_"))) alt ^ problem_suffix | 
| 51646 | 282 | |
| 283 | fun generate_casc_lbt_isa_files_for_theory ctxt thy format infer_policy type_enc | |
| 284 | dir_name = | |
| 285 | let | |
| 286 | val dir = Path.explode dir_name | |
| 287 | val _ = Isabelle_System.mkdir dir | |
| 288 | val file_order_path = ap dir file_order_name | |
| 289 | val _ = File.write file_order_path "" | |
| 290 | val problem_order_path = ap dir problem_order_name | |
| 291 | val _ = File.write problem_order_path "" | |
| 292 | val problem_dir = ap dir problem_name | |
| 293 | val _ = Isabelle_System.mkdir problem_dir | |
| 294 | val include_dir = ap problem_dir include_name | |
| 295 | val _ = Isabelle_System.mkdir include_dir | |
| 296 | val (prelude, groups) = | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 297 | problem_of_theory ctxt thy format infer_policy type_enc | 
| 51646 | 298 | |> split_last | 
| 299 | ||> (snd | |
| 300 | #> chop_maximal_groups (op = o pairself theory_name_of_fact) | |
| 301 | #> map (`(include_base_name_of_fact o hd))) | |
| 302 | fun write_prelude () = | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 303 | let val ss = lines_of_problem ctxt format prelude in | 
| 51646 | 304 | File.append file_order_path (prelude_base_name ^ "\n"); | 
| 305 | write_lines (ap include_dir prelude_name) ss | |
| 306 | end | |
| 307 | fun write_include_file (base_name, facts) = | |
| 308 | let | |
| 309 | val name = base_name ^ include_suffix | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 310 | val ss = lines_of_problem ctxt format [(factsN, facts)] | 
| 51646 | 311 | in | 
| 312 | File.append file_order_path (base_name ^ "\n"); | |
| 313 | write_lines (ap include_dir name) ss | |
| 314 | end | |
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 315 | fun write_problem_files _ _ _ [] = () | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 316 | | write_problem_files _ includes [] groups = | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 317 | write_problem_files 1 includes includes groups | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 318 | | write_problem_files n includes _ ((base_name, []) :: groups) = | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 319 | write_problem_files n (includes @ [include_line base_name]) [] groups | 
| 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 320 | | write_problem_files n includes seen | 
| 51646 | 321 | ((base_name, fact :: facts) :: groups) = | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 322 | let val fact_s = tptp_string_of_line format fact in | 
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 323 | (if should_generate_problem thy base_name fact then | 
| 51646 | 324 | let | 
| 325 | val (name, conj) = | |
| 326 | (case fact of | |
| 327 | Formula ((ident, alt), _, phi, _, _) => | |
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 328 | (problem_name_of base_name n (encode_meta alt), | 
| 51646 | 329 | Formula ((ident, alt), Conjecture, phi, NONE, []))) | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51652diff
changeset | 330 | val conj_s = tptp_string_of_line format conj | 
| 51646 | 331 | in | 
| 332 | File.append problem_order_path (name ^ "\n"); | |
| 333 | write_lines (ap problem_dir name) (seen @ [conj_s]) | |
| 334 | end | |
| 335 | else | |
| 336 | (); | |
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 337 | write_problem_files (n + 1) includes (seen @ [fact_s]) | 
| 51646 | 338 | ((base_name, facts) :: groups)) | 
| 339 | end | |
| 340 | val _ = write_prelude () | |
| 341 | val _ = app write_include_file groups | |
| 51652 
5ff01d585a8c
handle case clashes on Mac file system by encoding goal numbers
 blanchet parents: 
51650diff
changeset | 342 | val _ = write_problem_files 1 [include_line prelude_base_name] [] groups | 
| 51646 | 343 | in () end | 
| 344 | ||
| 42602 
a2db47fa015e
added TPTP exporter facility -- useful to do experiments with machine learning
 blanchet parents: diff
changeset | 345 | end; |