| author | ballarin | 
| Thu, 13 May 2010 14:47:15 +0200 | |
| changeset 36906 | 9eff24f4e5db | 
| parent 36568 | d495d2e1f0a6 | 
| child 36910 | dd5a31098f85 | 
| permissions | -rw-r--r-- | 
| 36376 | 1 | (* Title: HOL/Tools/ATP_Manager/atp_systems.ML | 
| 28592 | 2 | Author: Fabian Immler, TU Muenchen | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 28592 | 4 | |
| 36376 | 5 | Setup for supported ATPs. | 
| 28592 | 6 | *) | 
| 7 | ||
| 36376 | 8 | signature ATP_SYSTEMS = | 
| 28592 | 9 | sig | 
| 35867 | 10 | type prover = ATP_Manager.prover | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 11 | |
| 35867 | 12 | (* hooks for problem files *) | 
| 36376 | 13 | val dest_dir : string Config.T | 
| 35867 | 14 | val problem_prefix : string Config.T | 
| 15 | val measure_runtime : bool Config.T | |
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 16 | |
| 35867 | 17 | val refresh_systems_on_tptp : unit -> unit | 
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 18 | val default_atps_param_value : unit -> string | 
| 35867 | 19 | val setup : theory -> theory | 
| 28592 | 20 | end; | 
| 21 | ||
| 36376 | 22 | structure ATP_Systems : ATP_SYSTEMS = | 
| 28592 | 23 | struct | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 24 | |
| 36187 | 25 | open Sledgehammer_Util | 
| 35969 | 26 | open Sledgehammer_Fact_Preprocessor | 
| 35865 | 27 | open Sledgehammer_HOL_Clause | 
| 28 | open Sledgehammer_Fact_Filter | |
| 29 | open Sledgehammer_Proof_Reconstruct | |
| 35867 | 30 | open ATP_Manager | 
| 35826 | 31 | |
| 36376 | 32 | (** generic ATP **) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 33 | |
| 32944 
ecc0705174c2
clarified File.platform_path vs. File.shell_path;
 wenzelm parents: 
32942diff
changeset | 34 | (* external problem files *) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 35 | |
| 36376 | 36 | val (dest_dir, dest_dir_setup) = Attrib.config_string "atp_dest_dir" (K ""); | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 37 | (*Empty string means create files in Isabelle's temporary files directory.*) | 
| 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 38 | |
| 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 39 | val (problem_prefix, problem_prefix_setup) = | 
| 36001 | 40 | Attrib.config_string "atp_problem_prefix" (K "prob"); | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 41 | |
| 33247 | 42 | val (measure_runtime, measure_runtime_setup) = | 
| 36001 | 43 | Attrib.config_bool "atp_measure_runtime" (K false); | 
| 33247 | 44 | |
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 45 | |
| 35867 | 46 | (* prover configuration *) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 47 | |
| 32941 
72d48e333b77
eliminated extraneous wrapping of public records;
 wenzelm parents: 
32936diff
changeset | 48 | type prover_config = | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 49 |   {home: string,
 | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 50 | executable: string, | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 51 | arguments: Time.time -> string, | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 52 | proof_delims: (string * string) list, | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 53 | known_failures: (failure * string) list, | 
| 36382 | 54 | max_axiom_clauses: int, | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 55 | prefers_theory_relevant: bool}; | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 56 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 57 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 58 | (* basic template *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 59 | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 60 | val remotify = prefix "remote_" | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 61 | |
| 32458 
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
 boehmes parents: 
32451diff
changeset | 62 | fun with_path cleanup after f path = | 
| 
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
 boehmes parents: 
32451diff
changeset | 63 | Exn.capture f path | 
| 
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
 boehmes parents: 
32451diff
changeset | 64 | |> tap (fn _ => cleanup path) | 
| 
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
 boehmes parents: 
32451diff
changeset | 65 | |> Exn.release | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 66 | |> tap (after path) | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 67 | |
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 68 | (* Splits by the first possible of a list of delimiters. *) | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 69 | fun extract_proof delims output = | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 70 | case pairself (find_first (fn s => String.isSubstring s output)) | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 71 | (ListPair.unzip delims) of | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 72 | (SOME begin_delim, SOME end_delim) => | 
| 36549 | 73 | (output |> first_field begin_delim |> the |> snd | 
| 74 | |> first_field end_delim |> the |> fst | |
| 75 | |> first_field "\n" |> the |> snd | |
| 76 | handle Option.Option => "") | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 77 | | _ => "" | 
| 32458 
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
 boehmes parents: 
32451diff
changeset | 78 | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 79 | fun extract_proof_and_outcome res_code proof_delims known_failures output = | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 80 | case map_filter (fn (failure, pattern) => | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 81 | if String.isSubstring pattern output then SOME failure | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 82 | else NONE) known_failures of | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 83 | [] => (case extract_proof proof_delims output of | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 84 |              "" => ("", SOME UnknownError)
 | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 85 | | proof => if res_code = 0 then (proof, NONE) | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 86 |                       else ("", SOME UnknownError))
 | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 87 |   | (failure :: _) => ("", SOME failure)
 | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 88 | |
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 89 | fun string_for_failure Unprovable = "The ATP problem is unprovable." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 90 | | string_for_failure TimedOut = "Timed out." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 91 | | string_for_failure OutOfResources = "The ATP ran out of resources." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 92 | | string_for_failure OldSpass = | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36382diff
changeset | 93 | (* FIXME: Change the error message below to point to the Isabelle download | 
| 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36382diff
changeset | 94 | page once the package is there (around the Isabelle2010 release). *) | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 95 | "Warning: Sledgehammer requires a more recent version of SPASS with \ | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 96 | \support for the TPTP syntax. To install it, download and untar the \ | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 97 | \package \"http://isabelle.in.tum.de/~blanchet/spass-3.7.tgz\" and add the \ | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 98 | \\"spass-3.7\" directory's full path to \"" ^ | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 99 | Path.implode (Path.expand (Path.appends | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 100 | (Path.variable "ISABELLE_HOME_USER" :: | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 101 | map Path.basic ["etc", "components"]))) ^ | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 102 | "\" on a line of its own." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 103 | | string_for_failure MalformedOutput = "Error: The ATP output is malformed." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 104 | | string_for_failure UnknownError = "Error: An unknown ATP error occurred." | 
| 35865 | 105 | |
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 106 | fun shape_of_clauses _ [] = [] | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 107 | | shape_of_clauses j ([] :: clauses) = [] :: shape_of_clauses j clauses | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 108 | | shape_of_clauses j ((lit :: lits) :: clauses) = | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 109 | let val shape = shape_of_clauses (j + 1) (lits :: clauses) in | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 110 | (j :: hd shape) :: tl shape | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 111 | end | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 112 | |
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 113 | fun generic_prover overlord get_facts prepare write_file home executable args | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 114 | proof_delims known_failures name | 
| 36489 | 115 |         ({debug, full_types, explicit_apply, isar_proof, shrink_factor, ...}
 | 
| 116 | : params) minimize_command | |
| 35969 | 117 |         ({subgoal, goal, relevance_override, axiom_clauses, filtered_clauses}
 | 
| 118 | : problem) = | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 119 | let | 
| 31750 | 120 | (* get clauses and prepare them for writing *) | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 121 | val (ctxt, (chain_ths, th)) = goal; | 
| 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 122 | val thy = ProofContext.theory_of ctxt; | 
| 35865 | 123 | val chain_ths = map (Thm.put_name_hint chained_hint) chain_ths; | 
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 124 | val goal_clss = #1 (neg_conjecture_clauses ctxt th subgoal) | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 125 | val goal_cls = List.concat goal_clss | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31751diff
changeset | 126 | val the_filtered_clauses = | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 127 | (case filtered_clauses of | 
| 35969 | 128 | NONE => get_facts relevance_override goal goal_cls | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 129 | | SOME fcls => fcls); | 
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
31368diff
changeset | 130 | val the_axiom_clauses = | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 131 | (case axiom_clauses of | 
| 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 132 | NONE => the_filtered_clauses | 
| 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 133 | | SOME axcls => axcls); | 
| 35969 | 134 | val (internal_thm_names, clauses) = | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 135 | prepare goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy; | 
| 31750 | 136 | |
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 137 | (* path to unique problem file *) | 
| 36376 | 138 | val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER" | 
| 139 | else Config.get ctxt dest_dir; | |
| 140 | val the_problem_prefix = Config.get ctxt problem_prefix; | |
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 141 | fun prob_pathname nr = | 
| 36143 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 blanchet parents: 
36142diff
changeset | 142 | let | 
| 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 blanchet parents: 
36142diff
changeset | 143 | val probfile = | 
| 36568 
d495d2e1f0a6
in "overlord" mode: ignore problem prefix specified in the .thy file
 blanchet parents: 
36552diff
changeset | 144 | Path.basic ((if overlord then "prob_" ^ name | 
| 
d495d2e1f0a6
in "overlord" mode: ignore problem prefix specified in the .thy file
 blanchet parents: 
36552diff
changeset | 145 | else the_problem_prefix ^ serial_string ()) | 
| 36143 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 blanchet parents: 
36142diff
changeset | 146 | ^ "_" ^ string_of_int nr) | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 147 | in | 
| 36376 | 148 | if the_dest_dir = "" then File.tmp_path probfile | 
| 149 | else if File.exists (Path.explode the_dest_dir) | |
| 150 | then Path.append (Path.explode the_dest_dir) probfile | |
| 151 |         else error ("No such directory: " ^ the_dest_dir ^ ".")
 | |
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 152 | end; | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 153 | |
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 154 | val command = Path.explode (home ^ "/" ^ executable) | 
| 31750 | 155 | (* write out problem file and call prover *) | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 156 | fun command_line probfile = | 
| 36284 
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
 blanchet parents: 
36283diff
changeset | 157 | (if Config.get ctxt measure_runtime then | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 158 |          "TIMEFORMAT='%3U'; { time " ^
 | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 159 | space_implode " " [File.shell_path command, args, | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 160 | File.shell_path probfile] ^ " ; } 2>&1" | 
| 36284 
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
 blanchet parents: 
36283diff
changeset | 161 | else | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 162 | space_implode " " ["exec", File.shell_path command, args, | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 163 | File.shell_path probfile, "2>&1"]) ^ | 
| 36284 
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
 blanchet parents: 
36283diff
changeset | 164 | (if overlord then | 
| 36286 
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
 blanchet parents: 
36284diff
changeset | 165 | " | sed 's/,/, /g' \ | 
| 36485 | 166 | \| sed 's/\\([^!=<]\\)\\([=|]\\)\\([^=>]\\)/\\1 \\2 \\3/g' \ | 
| 36286 
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
 blanchet parents: 
36284diff
changeset | 167 | \| sed 's/ / /g' | sed 's/| |/||/g' \ | 
| 
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
 blanchet parents: 
36284diff
changeset | 168 | \| sed 's/ = = =/===/g' \ | 
| 
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
 blanchet parents: 
36284diff
changeset | 169 | \| sed 's/= = /== /g'" | 
| 36284 
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
 blanchet parents: 
36283diff
changeset | 170 | else | 
| 
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
 blanchet parents: 
36283diff
changeset | 171 | "") | 
| 32510 | 172 | fun split_time s = | 
| 173 | let | |
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 174 | val split = String.tokens (fn c => str c = "\n"); | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 175 | val (output, t) = s |> split |> split_last |> apfst cat_lines; | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 176 | fun as_num f = f >> (fst o read_int); | 
| 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 177 | val num = as_num (Scan.many1 Symbol.is_ascii_digit); | 
| 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 178 | val digit = Scan.one Symbol.is_ascii_digit; | 
| 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 179 | val num3 = as_num (digit ::: digit ::: (digit >> single)); | 
| 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 180 | val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b); | 
| 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 181 | val as_time = the_default 0 o Scan.read Symbol.stopper time o explode; | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 182 | in (output, as_time t) end; | 
| 33247 | 183 | fun split_time' s = | 
| 184 | if Config.get ctxt measure_runtime then split_time s else (s, 0) | |
| 32458 
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
 boehmes parents: 
32451diff
changeset | 185 | fun run_on probfile = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 186 | if File.exists command then | 
| 36235 
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
 blanchet parents: 
36231diff
changeset | 187 | write_file full_types explicit_apply probfile clauses | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 188 | |> pair (apfst split_time' (bash_output (command_line probfile))) | 
| 36382 | 189 | else | 
| 190 |         error ("Bad executable: " ^ Path.implode command ^ ".");
 | |
| 28592 | 191 | |
| 36167 
c1a35be8e476
make Sledgehammer's output more debugging friendly
 blanchet parents: 
36143diff
changeset | 192 | (* If the problem file has not been exported, remove it; otherwise, export | 
| 
c1a35be8e476
make Sledgehammer's output more debugging friendly
 blanchet parents: 
36143diff
changeset | 193 | the proof file too. *) | 
| 36376 | 194 | fun cleanup probfile = | 
| 195 | if the_dest_dir = "" then try File.rm probfile else NONE | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 196 | fun export probfile (((output, _), _), _) = | 
| 36376 | 197 | if the_dest_dir = "" then | 
| 36187 | 198 | () | 
| 199 | else | |
| 200 | File.write (Path.explode (Path.implode probfile ^ "_proof")) | |
| 36282 
9a7c5b86a105
generate command-line in addition to timestamp in ATP output file, for debugging purposes
 blanchet parents: 
36281diff
changeset | 201 | ((if overlord then | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 202 | "% " ^ command_line probfile ^ "\n% " ^ timestamp () ^ | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 203 | "\n" | 
| 36282 
9a7c5b86a105
generate command-line in addition to timestamp in ATP output file, for debugging purposes
 blanchet parents: 
36281diff
changeset | 204 | else | 
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 205 | "") ^ output) | 
| 32257 
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
 wenzelm parents: 
32091diff
changeset | 206 | |
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 207 | val (((output, atp_run_time_in_msecs), res_code), | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 208 | (pool, conjecture_offset)) = | 
| 35570 
0e30eef52d85
basic simplification of external_prover signature;
 wenzelm parents: 
35010diff
changeset | 209 | with_path cleanup export run_on (prob_pathname subgoal); | 
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 210 | val conjecture_shape = shape_of_clauses (conjecture_offset + 1) goal_clss | 
| 36167 
c1a35be8e476
make Sledgehammer's output more debugging friendly
 blanchet parents: 
36143diff
changeset | 211 | (* Check for success and print out some information on failure. *) | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 212 | val (proof, outcome) = | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 213 | extract_proof_and_outcome res_code proof_delims known_failures output | 
| 35969 | 214 | val (message, relevant_thm_names) = | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 215 | case outcome of | 
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 216 | NONE => | 
| 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 217 | proof_text isar_proof | 
| 36489 | 218 | (pool, debug, shrink_factor, ctxt, conjecture_shape) | 
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 219 | (minimize_command, proof, internal_thm_names, th, subgoal) | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 220 | | SOME failure => (string_for_failure failure ^ "\n", []) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 221 | in | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36382diff
changeset | 222 |     {outcome = outcome, message = message, pool = pool,
 | 
| 35969 | 223 | relevant_thm_names = relevant_thm_names, | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 224 | atp_run_time_in_msecs = atp_run_time_in_msecs, output = output, | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 225 | proof = proof, internal_thm_names = internal_thm_names, | 
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 226 | conjecture_shape = conjecture_shape, | 
| 35969 | 227 | filtered_clauses = the_filtered_clauses} | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 228 | end; | 
| 28592 | 229 | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 230 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 231 | (* generic TPTP-based provers *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 232 | |
| 35865 | 233 | fun generic_tptp_prover | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 234 |         (name, {home, executable, arguments, proof_delims, known_failures,
 | 
| 36382 | 235 | max_axiom_clauses, prefers_theory_relevant}) | 
| 36264 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 236 |         (params as {debug, overlord, respect_no_atp, relevance_threshold,
 | 
| 36473 
8a5c99a1c965
remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
 blanchet parents: 
36400diff
changeset | 237 | convergence, theory_relevant, follow_defs, isar_proof, ...}) | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36265diff
changeset | 238 | minimize_command timeout = | 
| 36143 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 blanchet parents: 
36142diff
changeset | 239 | generic_prover overlord | 
| 36058 
8256d5a185bd
added "respect_no_atp" and "convergence" options to Sledgehammer;
 blanchet parents: 
36001diff
changeset | 240 | (get_relevant_facts respect_no_atp relevance_threshold convergence | 
| 36473 
8a5c99a1c965
remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
 blanchet parents: 
36400diff
changeset | 241 | follow_defs max_axiom_clauses | 
| 36220 
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
 blanchet parents: 
36219diff
changeset | 242 | (the_default prefers_theory_relevant theory_relevant)) | 
| 36473 
8a5c99a1c965
remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
 blanchet parents: 
36400diff
changeset | 243 | (prepare_clauses false) | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36382diff
changeset | 244 | (write_tptp_file (debug andalso overlord)) home | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 245 | executable (arguments timeout) proof_delims known_failures name params | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 246 | minimize_command | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 247 | |
| 35969 | 248 | fun tptp_prover name p = (name, generic_tptp_prover (name, p)); | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 249 | |
| 32941 
72d48e333b77
eliminated extraneous wrapping of public records;
 wenzelm parents: 
32936diff
changeset | 250 | |
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 251 | (** common provers **) | 
| 28592 | 252 | |
| 36382 | 253 | fun to_generous_secs time = (Time.toMilliseconds time + 999) div 1000 | 
| 36142 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36064diff
changeset | 254 | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 255 | (* Vampire *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 256 | |
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 257 | (* Vampire requires an explicit time limit. *) | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 258 | |
| 35969 | 259 | val vampire_config : prover_config = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 260 |   {home = getenv "VAMPIRE_HOME",
 | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 261 | executable = "vampire", | 
| 36382 | 262 | arguments = fn timeout => | 
| 263 | "--output_syntax tptp --mode casc -t " ^ | |
| 264 | string_of_int (to_generous_secs timeout), | |
| 36548 
a8a6d7172c8c
try out Vampire 11 and parse its output correctly;
 blanchet parents: 
36489diff
changeset | 265 | proof_delims = | 
| 
a8a6d7172c8c
try out Vampire 11 and parse its output correctly;
 blanchet parents: 
36489diff
changeset | 266 |      [("=========== Refutation ==========",
 | 
| 
a8a6d7172c8c
try out Vampire 11 and parse its output correctly;
 blanchet parents: 
36489diff
changeset | 267 | "======= End of refutation ======="), | 
| 
a8a6d7172c8c
try out Vampire 11 and parse its output correctly;
 blanchet parents: 
36489diff
changeset | 268 |       ("% SZS output start Refutation", "% SZS output end Refutation")],
 | 
| 36265 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 269 | known_failures = | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 270 | [(Unprovable, "Satisfiability detected"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 271 | (OutOfResources, "CANNOT PROVE"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 272 | (OutOfResources, "Refutation not found")], | 
| 36382 | 273 | max_axiom_clauses = 60, | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 274 | prefers_theory_relevant = false} | 
| 35969 | 275 | val vampire = tptp_prover "vampire" vampire_config | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 276 | |
| 28592 | 277 | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 278 | (* E prover *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 279 | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 280 | val tstp_proof_delims = | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 281 |   ("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation")
 | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 282 | |
| 35969 | 283 | val e_config : prover_config = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 284 |   {home = getenv "E_HOME",
 | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 285 | executable = "eproof", | 
| 36382 | 286 | arguments = fn timeout => | 
| 287 | "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ | |
| 288 | string_of_int (to_generous_secs timeout), | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 289 | proof_delims = [tstp_proof_delims], | 
| 36265 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 290 | known_failures = | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 291 | [(Unprovable, "SZS status: Satisfiable"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 292 | (Unprovable, "SZS status Satisfiable"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 293 | (TimedOut, "Failure: Resource limit exceeded (time)"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 294 | (TimedOut, "time limit exceeded"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 295 | (OutOfResources, | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 296 | "# Cannot determine problem status within resource limit"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 297 | (OutOfResources, "SZS status: ResourceOut"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 298 | (OutOfResources, "SZS status ResourceOut")], | 
| 36382 | 299 | max_axiom_clauses = 100, | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 300 | prefers_theory_relevant = false} | 
| 35969 | 301 | val e = tptp_prover "e" e_config | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 302 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 303 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 304 | (* SPASS *) | 
| 28592 | 305 | |
| 35865 | 306 | fun generic_dfg_prover | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 307 |         (name, {home, executable, arguments, proof_delims, known_failures,
 | 
| 36382 | 308 | max_axiom_clauses, prefers_theory_relevant}) | 
| 36143 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 blanchet parents: 
36142diff
changeset | 309 |         (params as {overlord, respect_no_atp, relevance_threshold, convergence,
 | 
| 36473 
8a5c99a1c965
remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
 blanchet parents: 
36400diff
changeset | 310 | theory_relevant, follow_defs, ...}) | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36265diff
changeset | 311 | minimize_command timeout = | 
| 36143 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 blanchet parents: 
36142diff
changeset | 312 | generic_prover overlord | 
| 36058 
8256d5a185bd
added "respect_no_atp" and "convergence" options to Sledgehammer;
 blanchet parents: 
36001diff
changeset | 313 | (get_relevant_facts respect_no_atp relevance_threshold convergence | 
| 36473 
8a5c99a1c965
remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
 blanchet parents: 
36400diff
changeset | 314 | follow_defs max_axiom_clauses | 
| 36220 
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
 blanchet parents: 
36219diff
changeset | 315 | (the_default prefers_theory_relevant theory_relevant)) | 
| 36473 
8a5c99a1c965
remove "higher_order" option from Sledgehammer -- the "smart" default is good enough
 blanchet parents: 
36400diff
changeset | 316 | (prepare_clauses true) write_dfg_file home executable | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 317 | (arguments timeout) proof_delims known_failures name params | 
| 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 318 | minimize_command | 
| 32869 
159309603edc
recovered support for Spass: re-enabled writing problems in DFG format
 boehmes parents: 
32864diff
changeset | 319 | |
| 36264 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 320 | fun dfg_prover name p = (name, generic_dfg_prover (name, p)) | 
| 32869 
159309603edc
recovered support for Spass: re-enabled writing problems in DFG format
 boehmes parents: 
32864diff
changeset | 321 | |
| 36219 
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
 blanchet parents: 
36190diff
changeset | 322 | (* The "-VarWeight=3" option helps the higher-order problems, probably by | 
| 
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
 blanchet parents: 
36190diff
changeset | 323 | counteracting the presence of "hAPP". *) | 
| 36059 | 324 | val spass_config : prover_config = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 325 |   {home = getenv "SPASS_HOME",
 | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 326 | executable = "SPASS", | 
| 36382 | 327 | arguments = fn timeout => | 
| 328 | "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \ | |
| 329 | \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout), | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 330 |    proof_delims = [("Here is a proof", "Formulae used in the proof")],
 | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 331 | known_failures = | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 332 | [(Unprovable, "SPASS beiseite: Completion found"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 333 | (TimedOut, "SPASS beiseite: Ran out of time"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 334 | (OutOfResources, "SPASS beiseite: Maximal number of loops exceeded")], | 
| 36382 | 335 | max_axiom_clauses = 40, | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 336 | prefers_theory_relevant = true} | 
| 36264 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 337 | val spass = dfg_prover "spass" spass_config | 
| 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 338 | |
| 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 339 | (* SPASS 3.7 supports both the DFG and the TPTP syntax, whereas SPASS 3.0 | 
| 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 340 | supports only the DFG syntax. As soon as all Isabelle repository/snapshot | 
| 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 341 | users have upgraded to 3.7, we can kill "spass" (and all DFG support in | 
| 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 342 | Sledgehammer) and rename "spass_tptp" "spass". *) | 
| 35969 | 343 | |
| 36264 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 344 | val spass_tptp_config = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 345 |   {home = #home spass_config,
 | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 346 | executable = #executable spass_config, | 
| 36264 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 347 | arguments = prefix "-TPTP " o #arguments spass_config, | 
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 348 | proof_delims = #proof_delims spass_config, | 
| 36265 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 349 | known_failures = | 
| 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 350 | #known_failures spass_config @ | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 351 | [(OldSpass, "unrecognized option `-TPTP'"), | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 352 | (OldSpass, "Unrecognized option TPTP")], | 
| 36382 | 353 | max_axiom_clauses = #max_axiom_clauses spass_config, | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 354 | prefers_theory_relevant = #prefers_theory_relevant spass_config} | 
| 36264 
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
 blanchet parents: 
36235diff
changeset | 355 | val spass_tptp = tptp_prover "spass_tptp" spass_tptp_config | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 356 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 357 | (* remote prover invocation via SystemOnTPTP *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 358 | |
| 36376 | 359 | val systems = Synchronized.var "atp_systems" ([]: string list); | 
| 31835 | 360 | |
| 361 | fun get_systems () = | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 362 | case bash_output "\"$ISABELLE_ATP_MANAGER/SystemOnTPTP\" -w" of | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 363 | (answer, 0) => split_lines answer | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 364 | | (answer, _) => | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 365 |     error ("Failed to get available systems at SystemOnTPTP:\n" ^ answer)
 | 
| 31835 | 366 | |
| 35867 | 367 | fun refresh_systems_on_tptp () = | 
| 368 | Synchronized.change systems (fn _ => get_systems ()); | |
| 31835 | 369 | |
| 370 | fun get_system prefix = Synchronized.change_result systems (fn systems => | |
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 371 | (if null systems then get_systems () else systems) | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 372 | |> `(find_first (String.isPrefix prefix))); | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 373 | |
| 32948 | 374 | fun the_system prefix = | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 375 | (case get_system prefix of | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 376 |     NONE => error ("System " ^ quote prefix ^
 | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 377 | " not available at SystemOnTPTP.") | 
| 32942 
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
 wenzelm parents: 
32941diff
changeset | 378 | | SOME sys => sys); | 
| 31835 | 379 | |
| 36265 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 380 | val remote_known_failures = | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 381 | [(TimedOut, "says Timeout"), | 
| 36377 
b3dce4c715d0
now rename the file "atp_wrapper.ML" to "atp_systems.ML" + fix typo in "SystemOnTPTP" script
 blanchet parents: 
36376diff
changeset | 382 | (MalformedOutput, "Remote script could not extract proof")] | 
| 35865 | 383 | |
| 36382 | 384 | fun remote_prover_config atp_prefix args | 
| 385 |         ({proof_delims, known_failures, max_axiom_clauses,
 | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 386 | prefers_theory_relevant, ...} : prover_config) : prover_config = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 387 |   {home = getenv "ISABELLE_ATP_MANAGER",
 | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 388 | executable = "SystemOnTPTP", | 
| 36382 | 389 | arguments = fn timeout => | 
| 390 | args ^ " -t " ^ string_of_int (to_generous_secs timeout) ^ " -s " ^ | |
| 391 | the_system atp_prefix, | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 392 | proof_delims = insert (op =) tstp_proof_delims proof_delims, | 
| 36265 
41c9e755e552
distinguish between the different ATP errors in the user interface;
 blanchet parents: 
36264diff
changeset | 393 | known_failures = remote_known_failures @ known_failures, | 
| 36382 | 394 | max_axiom_clauses = max_axiom_clauses, | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 395 | prefers_theory_relevant = prefers_theory_relevant} | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 396 | |
| 35969 | 397 | val remote_vampire = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 398 | tptp_prover (remotify (fst vampire)) | 
| 36552 
2c042d86c711
back to Vampire 9 -- Vampire 11 sometimes outputs really weird proofs
 blanchet parents: 
36549diff
changeset | 399 | (remote_prover_config "Vampire---9" "" vampire_config) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 400 | |
| 35969 | 401 | val remote_e = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 402 | tptp_prover (remotify (fst e)) | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 403 | (remote_prover_config "EP---" "" e_config) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32740diff
changeset | 404 | |
| 35969 | 405 | val remote_spass = | 
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 406 | tptp_prover (remotify (fst spass)) | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 407 | (remote_prover_config "SPASS---" "-x" spass_config) | 
| 28592 | 408 | |
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 409 | fun maybe_remote (name, _) ({home, ...} : prover_config) =
 | 
| 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 410 | name |> home = "" ? remotify | 
| 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 411 | |
| 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 412 | fun default_atps_param_value () = | 
| 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 413 | space_implode " " [maybe_remote e e_config, maybe_remote spass spass_config, | 
| 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 414 | remotify (fst vampire)] | 
| 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 415 | |
| 36289 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 416 | val provers = | 
| 
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
 blanchet parents: 
36287diff
changeset | 417 | [spass, spass_tptp, vampire, e, remote_vampire, remote_spass, remote_e] | 
| 35867 | 418 | val prover_setup = fold add_prover provers | 
| 419 | ||
| 420 | val setup = | |
| 36376 | 421 | dest_dir_setup | 
| 35867 | 422 | #> problem_prefix_setup | 
| 423 | #> measure_runtime_setup | |
| 36371 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 blanchet parents: 
36370diff
changeset | 424 | #> prover_setup | 
| 35867 | 425 | |
| 28592 | 426 | end; |