blanchet@41087
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_provers.ML
|
wenzelm@28477
|
2 |
Author: Fabian Immler, TU Muenchen
|
wenzelm@32996
|
3 |
Author: Makarius
|
blanchet@35969
|
4 |
Author: Jasmin Blanchette, TU Muenchen
|
wenzelm@28477
|
5 |
|
blanchet@41087
|
6 |
Generic prover abstraction for Sledgehammer.
|
wenzelm@28477
|
7 |
*)
|
wenzelm@28477
|
8 |
|
blanchet@41087
|
9 |
signature SLEDGEHAMMER_PROVERS =
|
wenzelm@28477
|
10 |
sig
|
blanchet@40181
|
11 |
type failure = ATP_Proof.failure
|
blanchet@38988
|
12 |
type locality = Sledgehammer_Filter.locality
|
blanchet@40070
|
13 |
type relevance_fudge = Sledgehammer_Filter.relevance_fudge
|
blanchet@40114
|
14 |
type translated_formula = Sledgehammer_ATP_Translate.translated_formula
|
blanchet@40068
|
15 |
type minimize_command = Sledgehammer_ATP_Reconstruct.minimize_command
|
blanchet@39493
|
16 |
|
blanchet@35969
|
17 |
type params =
|
blanchet@38982
|
18 |
{blocking: bool,
|
blanchet@38982
|
19 |
debug: bool,
|
blanchet@35969
|
20 |
verbose: bool,
|
blanchet@36143
|
21 |
overlord: bool,
|
blanchet@40059
|
22 |
provers: string list,
|
blanchet@35969
|
23 |
full_types: bool,
|
blanchet@36235
|
24 |
explicit_apply: bool,
|
blanchet@38745
|
25 |
relevance_thresholds: real * real,
|
blanchet@38744
|
26 |
max_relevant: int option,
|
blanchet@35969
|
27 |
isar_proof: bool,
|
blanchet@36924
|
28 |
isar_shrink_factor: int,
|
blanchet@38985
|
29 |
timeout: Time.time,
|
blanchet@38985
|
30 |
expect: string}
|
blanchet@39493
|
31 |
|
blanchet@40204
|
32 |
datatype fact =
|
blanchet@40114
|
33 |
Untranslated of (string * locality) * thm |
|
blanchet@40114
|
34 |
Translated of term * ((string * locality) * translated_formula) option
|
blanchet@40061
|
35 |
|
blanchet@40061
|
36 |
type prover_problem =
|
blanchet@39318
|
37 |
{state: Proof.state,
|
blanchet@38998
|
38 |
goal: thm,
|
blanchet@38998
|
39 |
subgoal: int,
|
blanchet@40065
|
40 |
subgoal_count: int,
|
blanchet@40983
|
41 |
facts: fact list}
|
blanchet@39493
|
42 |
|
blanchet@40061
|
43 |
type prover_result =
|
blanchet@36370
|
44 |
{outcome: failure option,
|
blanchet@40204
|
45 |
used_facts: (string * locality) list,
|
blanchet@40062
|
46 |
run_time_in_msecs: int option,
|
blanchet@40062
|
47 |
message: string}
|
blanchet@39493
|
48 |
|
blanchet@40061
|
49 |
type prover = params -> minimize_command -> prover_problem -> prover_result
|
blanchet@35867
|
50 |
|
blanchet@41087
|
51 |
val is_smt_prover : Proof.context -> string -> bool
|
blanchet@40941
|
52 |
val is_prover_available : Proof.context -> string -> bool
|
blanchet@40072
|
53 |
val is_prover_installed : Proof.context -> string -> bool
|
blanchet@40941
|
54 |
val default_max_relevant_for_prover : Proof.context -> string -> int
|
blanchet@40369
|
55 |
val is_built_in_const_for_prover :
|
blanchet@41066
|
56 |
Proof.context -> string -> string * typ -> term list -> bool
|
blanchet@41087
|
57 |
val atp_relevance_fudge : relevance_fudge
|
blanchet@41087
|
58 |
val smt_relevance_fudge : relevance_fudge
|
blanchet@40941
|
59 |
val relevance_fudge_for_prover : Proof.context -> string -> relevance_fudge
|
blanchet@38023
|
60 |
val dest_dir : string Config.T
|
blanchet@38023
|
61 |
val problem_prefix : string Config.T
|
blanchet@39003
|
62 |
val measure_run_time : bool Config.T
|
blanchet@40941
|
63 |
val available_provers : Proof.context -> unit
|
blanchet@40059
|
64 |
val kill_provers : unit -> unit
|
blanchet@40059
|
65 |
val running_provers : unit -> unit
|
blanchet@40059
|
66 |
val messages : int option -> unit
|
blanchet@40941
|
67 |
val get_prover : Proof.context -> bool -> string -> prover
|
blanchet@41087
|
68 |
val run_prover :
|
blanchet@41087
|
69 |
params -> bool -> (string -> minimize_command) -> bool -> prover_problem
|
blanchet@41087
|
70 |
-> string -> bool * Proof.state
|
blanchet@38023
|
71 |
val setup : theory -> theory
|
wenzelm@28477
|
72 |
end;
|
wenzelm@28477
|
73 |
|
blanchet@41087
|
74 |
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS =
|
wenzelm@28477
|
75 |
struct
|
wenzelm@28477
|
76 |
|
blanchet@38028
|
77 |
open ATP_Problem
|
blanchet@39491
|
78 |
open ATP_Proof
|
blanchet@38028
|
79 |
open ATP_Systems
|
blanchet@39494
|
80 |
open Metis_Translate
|
blanchet@38023
|
81 |
open Sledgehammer_Util
|
blanchet@38988
|
82 |
open Sledgehammer_Filter
|
blanchet@40068
|
83 |
open Sledgehammer_ATP_Translate
|
blanchet@40068
|
84 |
open Sledgehammer_ATP_Reconstruct
|
blanchet@37583
|
85 |
|
blanchet@37583
|
86 |
(** The Sledgehammer **)
|
blanchet@37583
|
87 |
|
blanchet@38102
|
88 |
(* Identifier to distinguish Sledgehammer from other tools using
|
blanchet@38102
|
89 |
"Async_Manager". *)
|
blanchet@37585
|
90 |
val das_Tool = "Sledgehammer"
|
blanchet@37585
|
91 |
|
blanchet@40941
|
92 |
fun is_smt_prover ctxt name =
|
blanchet@40941
|
93 |
let val smts = SMT_Solver.available_solvers_of ctxt in
|
blanchet@40941
|
94 |
case try (unprefix remote_prefix) name of
|
blanchet@40941
|
95 |
SOME suffix => member (op =) smts suffix andalso
|
blanchet@40941
|
96 |
SMT_Solver.is_remotely_available ctxt suffix
|
blanchet@40941
|
97 |
| NONE => member (op =) smts name
|
blanchet@40941
|
98 |
end
|
blanchet@40062
|
99 |
|
blanchet@40941
|
100 |
fun is_prover_available ctxt name =
|
blanchet@40941
|
101 |
let val thy = ProofContext.theory_of ctxt in
|
blanchet@40941
|
102 |
is_smt_prover ctxt name orelse member (op =) (available_atps thy) name
|
blanchet@40941
|
103 |
end
|
blanchet@40072
|
104 |
|
blanchet@40072
|
105 |
fun is_prover_installed ctxt name =
|
blanchet@40072
|
106 |
let val thy = ProofContext.theory_of ctxt in
|
blanchet@40941
|
107 |
if is_smt_prover ctxt name then
|
blanchet@40941
|
108 |
String.isPrefix remote_prefix name orelse
|
blanchet@40941
|
109 |
SMT_Solver.is_locally_installed ctxt name
|
blanchet@40941
|
110 |
else
|
blanchet@40941
|
111 |
is_atp_installed thy name
|
blanchet@40941
|
112 |
end
|
blanchet@40941
|
113 |
|
blanchet@40941
|
114 |
fun available_smt_solvers ctxt =
|
blanchet@40941
|
115 |
let val smts = SMT_Solver.available_solvers_of ctxt in
|
blanchet@40941
|
116 |
smts @ map (prefix remote_prefix)
|
blanchet@40941
|
117 |
(filter (SMT_Solver.is_remotely_available ctxt) smts)
|
blanchet@40072
|
118 |
end
|
blanchet@40072
|
119 |
|
blanchet@40941
|
120 |
fun default_max_relevant_for_prover ctxt name =
|
blanchet@40941
|
121 |
let val thy = ProofContext.theory_of ctxt in
|
blanchet@40982
|
122 |
if is_smt_prover ctxt name then
|
blanchet@40982
|
123 |
SMT_Solver.default_max_relevant ctxt
|
blanchet@40982
|
124 |
(perhaps (try (unprefix remote_prefix)) name)
|
blanchet@40982
|
125 |
else
|
blanchet@40982
|
126 |
#default_max_relevant (get_atp thy name)
|
blanchet@40941
|
127 |
end
|
blanchet@40063
|
128 |
|
blanchet@40071
|
129 |
(* These are typically simplified away by "Meson.presimplify". Equality is
|
blanchet@40071
|
130 |
handled specially via "fequal". *)
|
blanchet@40071
|
131 |
val atp_irrelevant_consts =
|
blanchet@40071
|
132 |
[@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
|
blanchet@40071
|
133 |
@{const_name HOL.eq}]
|
blanchet@40206
|
134 |
|
blanchet@41066
|
135 |
fun is_built_in_const_for_prover ctxt name (s, T) args =
|
blanchet@41066
|
136 |
if is_smt_prover ctxt name then SMT_Builtin.is_builtin_ext ctxt (s, T) args
|
blanchet@40369
|
137 |
else member (op =) atp_irrelevant_consts s
|
blanchet@40071
|
138 |
|
blanchet@40070
|
139 |
(* FUDGE *)
|
blanchet@40070
|
140 |
val atp_relevance_fudge =
|
blanchet@40070
|
141 |
{worse_irrel_freq = 100.0,
|
blanchet@40070
|
142 |
higher_order_irrel_weight = 1.05,
|
blanchet@40070
|
143 |
abs_rel_weight = 0.5,
|
blanchet@40070
|
144 |
abs_irrel_weight = 2.0,
|
blanchet@40070
|
145 |
skolem_irrel_weight = 0.75,
|
blanchet@40070
|
146 |
theory_const_rel_weight = 0.5,
|
blanchet@40070
|
147 |
theory_const_irrel_weight = 0.25,
|
blanchet@40070
|
148 |
intro_bonus = 0.15,
|
blanchet@40070
|
149 |
elim_bonus = 0.15,
|
blanchet@40070
|
150 |
simp_bonus = 0.15,
|
blanchet@40070
|
151 |
local_bonus = 0.55,
|
blanchet@40070
|
152 |
assum_bonus = 1.05,
|
blanchet@40070
|
153 |
chained_bonus = 1.5,
|
blanchet@40070
|
154 |
max_imperfect = 11.5,
|
blanchet@40070
|
155 |
max_imperfect_exp = 1.0,
|
blanchet@40070
|
156 |
threshold_divisor = 2.0,
|
blanchet@40070
|
157 |
ridiculous_threshold = 0.1}
|
blanchet@40070
|
158 |
|
blanchet@40071
|
159 |
(* FUDGE (FIXME) *)
|
blanchet@40070
|
160 |
val smt_relevance_fudge =
|
blanchet@40071
|
161 |
{worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
|
blanchet@40071
|
162 |
higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
|
blanchet@40071
|
163 |
abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
|
blanchet@40071
|
164 |
abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
|
blanchet@40071
|
165 |
skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge,
|
blanchet@40071
|
166 |
theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
|
blanchet@40071
|
167 |
theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
|
blanchet@40071
|
168 |
intro_bonus = #intro_bonus atp_relevance_fudge,
|
blanchet@40071
|
169 |
elim_bonus = #elim_bonus atp_relevance_fudge,
|
blanchet@40071
|
170 |
simp_bonus = #simp_bonus atp_relevance_fudge,
|
blanchet@40071
|
171 |
local_bonus = #local_bonus atp_relevance_fudge,
|
blanchet@40071
|
172 |
assum_bonus = #assum_bonus atp_relevance_fudge,
|
blanchet@40071
|
173 |
chained_bonus = #chained_bonus atp_relevance_fudge,
|
blanchet@40071
|
174 |
max_imperfect = #max_imperfect atp_relevance_fudge,
|
blanchet@40071
|
175 |
max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
|
blanchet@40071
|
176 |
threshold_divisor = #threshold_divisor atp_relevance_fudge,
|
blanchet@40071
|
177 |
ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
|
blanchet@40070
|
178 |
|
blanchet@40941
|
179 |
fun relevance_fudge_for_prover ctxt name =
|
blanchet@40941
|
180 |
if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
|
blanchet@40070
|
181 |
|
blanchet@40941
|
182 |
fun available_provers ctxt =
|
blanchet@40060
|
183 |
let
|
blanchet@40941
|
184 |
val thy = ProofContext.theory_of ctxt
|
blanchet@40060
|
185 |
val (remote_provers, local_provers) =
|
blanchet@40941
|
186 |
sort_strings (available_atps thy) @
|
blanchet@40941
|
187 |
sort_strings (available_smt_solvers ctxt)
|
blanchet@40060
|
188 |
|> List.partition (String.isPrefix remote_prefix)
|
blanchet@40060
|
189 |
in
|
blanchet@40205
|
190 |
Output.urgent_message ("Available provers: " ^
|
blanchet@40205
|
191 |
commas (local_provers @ remote_provers) ^ ".")
|
blanchet@40060
|
192 |
end
|
blanchet@35969
|
193 |
|
blanchet@40059
|
194 |
fun kill_provers () = Async_Manager.kill_threads das_Tool "provers"
|
blanchet@40059
|
195 |
fun running_provers () = Async_Manager.running_threads das_Tool "provers"
|
blanchet@40059
|
196 |
val messages = Async_Manager.thread_messages das_Tool "prover"
|
blanchet@40059
|
197 |
|
blanchet@40059
|
198 |
(** problems, results, ATPs, etc. **)
|
blanchet@35969
|
199 |
|
blanchet@35969
|
200 |
type params =
|
blanchet@38982
|
201 |
{blocking: bool,
|
blanchet@38982
|
202 |
debug: bool,
|
blanchet@35969
|
203 |
verbose: bool,
|
blanchet@36143
|
204 |
overlord: bool,
|
blanchet@40059
|
205 |
provers: string list,
|
blanchet@35969
|
206 |
full_types: bool,
|
blanchet@36235
|
207 |
explicit_apply: bool,
|
blanchet@38745
|
208 |
relevance_thresholds: real * real,
|
blanchet@38744
|
209 |
max_relevant: int option,
|
blanchet@35969
|
210 |
isar_proof: bool,
|
blanchet@36924
|
211 |
isar_shrink_factor: int,
|
blanchet@38985
|
212 |
timeout: Time.time,
|
blanchet@38985
|
213 |
expect: string}
|
blanchet@35867
|
214 |
|
blanchet@40204
|
215 |
datatype fact =
|
blanchet@40114
|
216 |
Untranslated of (string * locality) * thm |
|
blanchet@40114
|
217 |
Translated of term * ((string * locality) * translated_formula) option
|
blanchet@40061
|
218 |
|
blanchet@40061
|
219 |
type prover_problem =
|
blanchet@39318
|
220 |
{state: Proof.state,
|
blanchet@38998
|
221 |
goal: thm,
|
blanchet@38998
|
222 |
subgoal: int,
|
blanchet@40065
|
223 |
subgoal_count: int,
|
blanchet@40983
|
224 |
facts: fact list}
|
blanchet@35867
|
225 |
|
blanchet@40061
|
226 |
type prover_result =
|
blanchet@36370
|
227 |
{outcome: failure option,
|
blanchet@35969
|
228 |
message: string,
|
blanchet@40204
|
229 |
used_facts: (string * locality) list,
|
blanchet@40062
|
230 |
run_time_in_msecs: int option}
|
blanchet@35867
|
231 |
|
blanchet@40061
|
232 |
type prover = params -> minimize_command -> prover_problem -> prover_result
|
blanchet@35867
|
233 |
|
blanchet@38023
|
234 |
(* configuration attributes *)
|
blanchet@38023
|
235 |
|
blanchet@38991
|
236 |
val (dest_dir, dest_dir_setup) =
|
blanchet@39003
|
237 |
Attrib.config_string "sledgehammer_dest_dir" (K "")
|
blanchet@38991
|
238 |
(* Empty string means create files in Isabelle's temporary files directory. *)
|
blanchet@38023
|
239 |
|
blanchet@38023
|
240 |
val (problem_prefix, problem_prefix_setup) =
|
blanchet@39003
|
241 |
Attrib.config_string "sledgehammer_problem_prefix" (K "prob")
|
blanchet@38023
|
242 |
|
blanchet@39003
|
243 |
val (measure_run_time, measure_run_time_setup) =
|
blanchet@39003
|
244 |
Attrib.config_bool "sledgehammer_measure_run_time" (K false)
|
wenzelm@28484
|
245 |
|
blanchet@38023
|
246 |
fun with_path cleanup after f path =
|
blanchet@38023
|
247 |
Exn.capture f path
|
blanchet@38023
|
248 |
|> tap (fn _ => cleanup path)
|
blanchet@38023
|
249 |
|> Exn.release
|
blanchet@38023
|
250 |
|> tap (after path)
|
blanchet@38023
|
251 |
|
blanchet@40204
|
252 |
fun prover_description ctxt ({blocking, verbose, ...} : params) name num_facts i
|
blanchet@40204
|
253 |
n goal =
|
blanchet@40060
|
254 |
quote name ^
|
blanchet@40060
|
255 |
(if verbose then
|
blanchet@40204
|
256 |
" with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts
|
blanchet@40060
|
257 |
else
|
blanchet@40060
|
258 |
"") ^
|
blanchet@40060
|
259 |
" on " ^ (if n = 1 then "goal" else "subgoal " ^ string_of_int i) ^ ":" ^
|
blanchet@40060
|
260 |
(if blocking then
|
blanchet@40060
|
261 |
""
|
blanchet@40060
|
262 |
else
|
blanchet@40060
|
263 |
"\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)))
|
blanchet@40060
|
264 |
|
blanchet@40060
|
265 |
fun proof_banner auto =
|
blanchet@40224
|
266 |
if auto then "Auto Sledgehammer found a proof" else "Try this command"
|
blanchet@40060
|
267 |
|
blanchet@40059
|
268 |
(* generic TPTP-based ATPs *)
|
blanchet@38023
|
269 |
|
blanchet@40114
|
270 |
fun dest_Untranslated (Untranslated p) = p
|
blanchet@40114
|
271 |
| dest_Untranslated (Translated _) = raise Fail "dest_Untranslated"
|
blanchet@40204
|
272 |
fun translated_fact ctxt (Untranslated p) = translate_fact ctxt p
|
blanchet@40204
|
273 |
| translated_fact _ (Translated p) = p
|
blanchet@40061
|
274 |
|
blanchet@40409
|
275 |
fun int_opt_add (SOME m) (SOME n) = SOME (m + n)
|
blanchet@40409
|
276 |
| int_opt_add _ _ = NONE
|
blanchet@40062
|
277 |
|
blanchet@39492
|
278 |
(* Important messages are important but not so important that users want to see
|
blanchet@39492
|
279 |
them each time. *)
|
blanchet@39493
|
280 |
val important_message_keep_factor = 0.1
|
blanchet@39492
|
281 |
|
blanchet@40063
|
282 |
fun run_atp auto atp_name
|
blanchet@38645
|
283 |
{exec, required_execs, arguments, has_incomplete_mode, proof_delims,
|
blanchet@40983
|
284 |
known_failures, explicit_forall, use_conjecture_for_hypotheses, ...}
|
blanchet@40983
|
285 |
({debug, verbose, overlord, full_types, explicit_apply, isar_proof,
|
blanchet@40983
|
286 |
isar_shrink_factor, timeout, ...} : params)
|
blanchet@40983
|
287 |
minimize_command ({state, goal, subgoal, facts, ...} : prover_problem) =
|
blanchet@38023
|
288 |
let
|
blanchet@39318
|
289 |
val ctxt = Proof.context_of state
|
blanchet@38998
|
290 |
val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
|
blanchet@40204
|
291 |
val facts =
|
blanchet@40983
|
292 |
facts |> map (translated_fact ctxt)
|
blanchet@40059
|
293 |
val dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
|
blanchet@40059
|
294 |
else Config.get ctxt dest_dir
|
blanchet@40059
|
295 |
val problem_prefix = Config.get ctxt problem_prefix
|
blanchet@40061
|
296 |
val problem_file_name =
|
blanchet@39003
|
297 |
Path.basic ((if overlord then "prob_" ^ atp_name
|
blanchet@40059
|
298 |
else problem_prefix ^ serial_string ())
|
blanchet@39003
|
299 |
^ "_" ^ string_of_int subgoal)
|
blanchet@40061
|
300 |
val problem_path_name =
|
blanchet@40059
|
301 |
if dest_dir = "" then
|
blanchet@40061
|
302 |
File.tmp_path problem_file_name
|
blanchet@40059
|
303 |
else if File.exists (Path.explode dest_dir) then
|
blanchet@40061
|
304 |
Path.append (Path.explode dest_dir) problem_file_name
|
blanchet@39003
|
305 |
else
|
blanchet@40059
|
306 |
error ("No such directory: " ^ quote dest_dir ^ ".")
|
blanchet@39003
|
307 |
val measure_run_time = verbose orelse Config.get ctxt measure_run_time
|
blanchet@38092
|
308 |
val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
|
blanchet@40059
|
309 |
(* write out problem file and call ATP *)
|
blanchet@38645
|
310 |
fun command_line complete timeout probfile =
|
blanchet@38023
|
311 |
let
|
blanchet@38023
|
312 |
val core = File.shell_path command ^ " " ^ arguments complete timeout ^
|
blanchet@38023
|
313 |
" " ^ File.shell_path probfile
|
blanchet@38023
|
314 |
in
|
blanchet@39010
|
315 |
(if measure_run_time then "TIMEFORMAT='%3R'; { time " ^ core ^ " ; }"
|
blanchet@38744
|
316 |
else "exec " ^ core) ^ " 2>&1"
|
blanchet@38023
|
317 |
end
|
blanchet@38023
|
318 |
fun split_time s =
|
blanchet@38023
|
319 |
let
|
blanchet@38023
|
320 |
val split = String.tokens (fn c => str c = "\n");
|
blanchet@38023
|
321 |
val (output, t) = s |> split |> split_last |> apfst cat_lines;
|
blanchet@38023
|
322 |
fun as_num f = f >> (fst o read_int);
|
blanchet@38023
|
323 |
val num = as_num (Scan.many1 Symbol.is_ascii_digit);
|
blanchet@38023
|
324 |
val digit = Scan.one Symbol.is_ascii_digit;
|
blanchet@38023
|
325 |
val num3 = as_num (digit ::: digit ::: (digit >> single));
|
blanchet@38023
|
326 |
val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b);
|
wenzelm@40627
|
327 |
val as_time = Scan.read Symbol.stopper time o raw_explode
|
blanchet@38023
|
328 |
in (output, as_time t) end;
|
blanchet@38023
|
329 |
fun run_on probfile =
|
blanchet@38092
|
330 |
case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of
|
blanchet@38032
|
331 |
(home_var, _) :: _ =>
|
blanchet@38023
|
332 |
error ("The environment variable " ^ quote home_var ^ " is not set.")
|
blanchet@38032
|
333 |
| [] =>
|
blanchet@38032
|
334 |
if File.exists command then
|
blanchet@38032
|
335 |
let
|
blanchet@39318
|
336 |
fun run complete timeout =
|
blanchet@38032
|
337 |
let
|
blanchet@38645
|
338 |
val command = command_line complete timeout probfile
|
blanchet@38032
|
339 |
val ((output, msecs), res_code) =
|
blanchet@38032
|
340 |
bash_output command
|
blanchet@38032
|
341 |
|>> (if overlord then
|
blanchet@38032
|
342 |
prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
|
blanchet@38032
|
343 |
else
|
blanchet@38032
|
344 |
I)
|
blanchet@40062
|
345 |
|>> (if measure_run_time then split_time else rpair NONE)
|
blanchet@39452
|
346 |
val (tstplike_proof, outcome) =
|
blanchet@39452
|
347 |
extract_tstplike_proof_and_outcome complete res_code
|
blanchet@39452
|
348 |
proof_delims known_failures output
|
blanchet@39452
|
349 |
in (output, msecs, tstplike_proof, outcome) end
|
blanchet@38032
|
350 |
val readable_names = debug andalso overlord
|
blanchet@40204
|
351 |
val (atp_problem, pool, conjecture_offset, fact_names) =
|
blanchet@40059
|
352 |
prepare_atp_problem ctxt readable_names explicit_forall full_types
|
blanchet@40204
|
353 |
explicit_apply hyp_ts concl_t facts
|
blanchet@39452
|
354 |
val ss = tptp_strings_for_atp_problem use_conjecture_for_hypotheses
|
blanchet@40059
|
355 |
atp_problem
|
blanchet@38631
|
356 |
val _ = File.write_list probfile ss
|
blanchet@38032
|
357 |
val conjecture_shape =
|
blanchet@38032
|
358 |
conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1
|
blanchet@38040
|
359 |
|> map single
|
blanchet@39318
|
360 |
val run_twice = has_incomplete_mode andalso not auto
|
blanchet@38645
|
361 |
val timer = Timer.startRealTimer ()
|
blanchet@38032
|
362 |
val result =
|
blanchet@39318
|
363 |
run false (if run_twice then
|
blanchet@39318
|
364 |
Time.fromMilliseconds
|
blanchet@38645
|
365 |
(2 * Time.toMilliseconds timeout div 3)
|
blanchet@39318
|
366 |
else
|
blanchet@39318
|
367 |
timeout)
|
blanchet@39318
|
368 |
|> run_twice
|
blanchet@38645
|
369 |
? (fn (_, msecs0, _, SOME _) =>
|
blanchet@39318
|
370 |
run true (Time.- (timeout, Timer.checkRealTimer timer))
|
blanchet@39452
|
371 |
|> (fn (output, msecs, tstplike_proof, outcome) =>
|
blanchet@40409
|
372 |
(output, int_opt_add msecs0 msecs, tstplike_proof,
|
blanchet@40409
|
373 |
outcome))
|
blanchet@38645
|
374 |
| result => result)
|
blanchet@40204
|
375 |
in ((pool, conjecture_shape, fact_names), result) end
|
blanchet@38032
|
376 |
else
|
blanchet@38032
|
377 |
error ("Bad executable: " ^ Path.implode command ^ ".")
|
blanchet@38023
|
378 |
|
blanchet@38023
|
379 |
(* If the problem file has not been exported, remove it; otherwise, export
|
blanchet@38023
|
380 |
the proof file too. *)
|
blanchet@38023
|
381 |
fun cleanup probfile =
|
blanchet@40059
|
382 |
if dest_dir = "" then try File.rm probfile else NONE
|
blanchet@38023
|
383 |
fun export probfile (_, (output, _, _, _)) =
|
blanchet@40059
|
384 |
if dest_dir = "" then
|
blanchet@38023
|
385 |
()
|
blanchet@38023
|
386 |
else
|
blanchet@38023
|
387 |
File.write (Path.explode (Path.implode probfile ^ "_proof")) output
|
blanchet@40204
|
388 |
val ((pool, conjecture_shape, fact_names),
|
blanchet@39452
|
389 |
(output, msecs, tstplike_proof, outcome)) =
|
blanchet@40061
|
390 |
with_path cleanup export run_on problem_path_name
|
blanchet@40204
|
391 |
val (conjecture_shape, fact_names) =
|
blanchet@40204
|
392 |
repair_conjecture_shape_and_fact_names output conjecture_shape fact_names
|
blanchet@39492
|
393 |
val important_message =
|
blanchet@40224
|
394 |
if not auto andalso random () <= important_message_keep_factor then
|
blanchet@39492
|
395 |
extract_important_message output
|
blanchet@39492
|
396 |
else
|
blanchet@39492
|
397 |
""
|
blanchet@40204
|
398 |
val (message, used_facts) =
|
blanchet@38023
|
399 |
case outcome of
|
blanchet@38023
|
400 |
NONE =>
|
blanchet@38023
|
401 |
proof_text isar_proof
|
blanchet@38023
|
402 |
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
|
blanchet@40060
|
403 |
(proof_banner auto, full_types, minimize_command, tstplike_proof,
|
blanchet@40204
|
404 |
fact_names, goal, subgoal)
|
blanchet@38744
|
405 |
|>> (fn message =>
|
blanchet@40341
|
406 |
message ^
|
blanchet@40341
|
407 |
(if verbose then
|
blanchet@40341
|
408 |
"\nATP real CPU time: " ^
|
blanchet@40341
|
409 |
string_from_time (Time.fromMilliseconds (the msecs)) ^ "."
|
blanchet@40341
|
410 |
else
|
blanchet@40341
|
411 |
"") ^
|
blanchet@39107
|
412 |
(if important_message <> "" then
|
blanchet@39107
|
413 |
"\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^
|
blanchet@39107
|
414 |
important_message
|
blanchet@39107
|
415 |
else
|
blanchet@39107
|
416 |
""))
|
blanchet@40669
|
417 |
| SOME failure => (string_for_failure "ATP" failure, [])
|
blanchet@38023
|
418 |
in
|
blanchet@40204
|
419 |
{outcome = outcome, message = message, used_facts = used_facts,
|
blanchet@40061
|
420 |
run_time_in_msecs = msecs}
|
blanchet@38023
|
421 |
end
|
blanchet@38023
|
422 |
|
blanchet@40669
|
423 |
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
|
blanchet@40669
|
424 |
these are sorted out properly in the SMT module, we have to interpret these
|
blanchet@40669
|
425 |
ourselves. *)
|
blanchet@40684
|
426 |
val remote_smt_failures =
|
blanchet@40684
|
427 |
[(22, CantConnect),
|
blanchet@40684
|
428 |
(127, NoPerl),
|
blanchet@40684
|
429 |
(2, NoLibwwwPerl)]
|
blanchet@40684
|
430 |
val z3_failures =
|
blanchet@40684
|
431 |
[(103, MalformedInput),
|
blanchet@40684
|
432 |
(110, MalformedInput)]
|
blanchet@40684
|
433 |
val unix_failures =
|
blanchet@40684
|
434 |
[(139, Crashed)]
|
blanchet@40684
|
435 |
val smt_failures = remote_smt_failures @ z3_failures @ unix_failures
|
blanchet@40555
|
436 |
|
boehmes@40424
|
437 |
fun failure_from_smt_failure (SMT_Failure.Counterexample _) = Unprovable
|
boehmes@40424
|
438 |
| failure_from_smt_failure SMT_Failure.Time_Out = TimedOut
|
boehmes@40561
|
439 |
| failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) =
|
blanchet@40684
|
440 |
(case AList.lookup (op =) smt_failures code of
|
blanchet@40684
|
441 |
SOME failure => failure
|
blanchet@40684
|
442 |
| NONE => UnknownError)
|
boehmes@40424
|
443 |
| failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
|
blanchet@40207
|
444 |
| failure_from_smt_failure _ = UnknownError
|
blanchet@40063
|
445 |
|
blanchet@40698
|
446 |
(* FUDGE *)
|
blanchet@40553
|
447 |
val smt_max_iter = 8
|
blanchet@40419
|
448 |
val smt_iter_fact_divisor = 2
|
blanchet@40419
|
449 |
val smt_iter_min_msecs = 5000
|
blanchet@40419
|
450 |
val smt_iter_timeout_divisor = 2
|
blanchet@40439
|
451 |
val smt_monomorph_limit = 4
|
blanchet@40409
|
452 |
|
blanchet@40723
|
453 |
fun smt_filter_loop ({verbose, timeout, ...} : params) remote state i =
|
blanchet@40409
|
454 |
let
|
blanchet@40553
|
455 |
val ctxt = Proof.context_of state
|
blanchet@40553
|
456 |
fun iter timeout iter_num outcome0 msecs_so_far facts =
|
blanchet@40553
|
457 |
let
|
blanchet@40553
|
458 |
val timer = Timer.startRealTimer ()
|
blanchet@40553
|
459 |
val ms = timeout |> Time.toMilliseconds
|
blanchet@40553
|
460 |
val iter_timeout =
|
blanchet@40553
|
461 |
if iter_num < smt_max_iter then
|
blanchet@40553
|
462 |
Int.min (ms, Int.max (smt_iter_min_msecs,
|
blanchet@40553
|
463 |
ms div smt_iter_timeout_divisor))
|
blanchet@40553
|
464 |
|> Time.fromMilliseconds
|
blanchet@40553
|
465 |
else
|
blanchet@40553
|
466 |
timeout
|
blanchet@40553
|
467 |
val num_facts = length facts
|
blanchet@40553
|
468 |
val _ =
|
blanchet@40553
|
469 |
if verbose then
|
blanchet@40553
|
470 |
"SMT iteration with " ^ string_of_int num_facts ^ " fact" ^
|
blanchet@40668
|
471 |
plural_s num_facts ^ " for " ^ string_from_time iter_timeout ^ "..."
|
blanchet@40553
|
472 |
|> Output.urgent_message
|
blanchet@40553
|
473 |
else
|
blanchet@40553
|
474 |
()
|
blanchet@40553
|
475 |
val {outcome, used_facts, run_time_in_msecs} =
|
blanchet@40979
|
476 |
SMT_Solver.smt_filter remote iter_timeout state facts i
|
blanchet@40553
|
477 |
val _ =
|
blanchet@40553
|
478 |
if verbose andalso is_some outcome then
|
blanchet@40553
|
479 |
"SMT outcome: " ^ SMT_Failure.string_of_failure ctxt (the outcome)
|
blanchet@40553
|
480 |
|> Output.urgent_message
|
blanchet@40553
|
481 |
else
|
blanchet@40553
|
482 |
()
|
blanchet@40553
|
483 |
val outcome0 = if is_none outcome0 then SOME outcome else outcome0
|
blanchet@40553
|
484 |
val msecs_so_far = int_opt_add run_time_in_msecs msecs_so_far
|
blanchet@40553
|
485 |
val too_many_facts_perhaps =
|
blanchet@40553
|
486 |
case outcome of
|
blanchet@40553
|
487 |
NONE => false
|
blanchet@40553
|
488 |
| SOME (SMT_Failure.Counterexample _) => false
|
blanchet@40553
|
489 |
| SOME SMT_Failure.Time_Out => iter_timeout <> timeout
|
boehmes@40561
|
490 |
| SOME (SMT_Failure.Abnormal_Termination code) =>
|
blanchet@40553
|
491 |
(if verbose then
|
blanchet@40555
|
492 |
"The SMT solver invoked with " ^ string_of_int num_facts ^
|
blanchet@40692
|
493 |
" fact" ^ plural_s num_facts ^ " terminated abnormally with \
|
blanchet@40692
|
494 |
\exit code " ^ string_of_int code ^ "."
|
blanchet@40692
|
495 |
|> warning
|
blanchet@40553
|
496 |
else
|
blanchet@40553
|
497 |
();
|
blanchet@40553
|
498 |
true (* kind of *))
|
blanchet@40553
|
499 |
| SOME SMT_Failure.Out_Of_Memory => true
|
blanchet@40553
|
500 |
| SOME _ => true
|
blanchet@40553
|
501 |
val timeout = Time.- (timeout, Timer.checkRealTimer timer)
|
blanchet@40553
|
502 |
in
|
blanchet@40553
|
503 |
if too_many_facts_perhaps andalso iter_num < smt_max_iter andalso
|
blanchet@40553
|
504 |
num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then
|
blanchet@40553
|
505 |
let val facts = take (num_facts div smt_iter_fact_divisor) facts in
|
blanchet@40553
|
506 |
iter timeout (iter_num + 1) outcome0 msecs_so_far facts
|
blanchet@40553
|
507 |
end
|
blanchet@40553
|
508 |
else
|
blanchet@40553
|
509 |
{outcome = if is_none outcome then NONE else the outcome0,
|
blanchet@40553
|
510 |
used_facts = used_facts, run_time_in_msecs = msecs_so_far}
|
blanchet@40409
|
511 |
end
|
blanchet@40553
|
512 |
in iter timeout 1 NONE (SOME 0) end
|
blanchet@40409
|
513 |
|
blanchet@40666
|
514 |
(* taken from "Mirabelle" and generalized *)
|
blanchet@40666
|
515 |
fun can_apply timeout tac state i =
|
blanchet@40666
|
516 |
let
|
blanchet@40666
|
517 |
val {context = ctxt, facts, goal} = Proof.goal state
|
blanchet@40666
|
518 |
val full_tac = Method.insert_tac facts i THEN tac ctxt i
|
blanchet@40666
|
519 |
in
|
blanchet@40666
|
520 |
case try (TimeLimit.timeLimit timeout (Seq.pull o full_tac)) goal of
|
blanchet@40666
|
521 |
SOME (SOME _) => true
|
blanchet@40666
|
522 |
| _ => false
|
blanchet@40666
|
523 |
end
|
blanchet@40666
|
524 |
|
blanchet@40666
|
525 |
val smt_metis_timeout = seconds 0.5
|
blanchet@40666
|
526 |
|
blanchet@40693
|
527 |
fun can_apply_metis debug state i ths =
|
blanchet@40693
|
528 |
can_apply smt_metis_timeout
|
blanchet@40693
|
529 |
(Config.put Metis_Tactics.verbose debug
|
blanchet@40693
|
530 |
#> (fn ctxt => Metis_Tactics.metis_tac ctxt ths)) state i
|
blanchet@40666
|
531 |
|
blanchet@40983
|
532 |
fun run_smt_solver auto name (params as {debug, ...}) minimize_command
|
blanchet@40983
|
533 |
({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
|
blanchet@36379
|
534 |
let
|
blanchet@40941
|
535 |
val (remote, suffix) =
|
blanchet@40941
|
536 |
case try (unprefix remote_prefix) name of
|
blanchet@40941
|
537 |
SOME suffix => (true, suffix)
|
blanchet@40941
|
538 |
| NONE => (false, name)
|
blanchet@40439
|
539 |
val repair_context =
|
blanchet@40941
|
540 |
Context.proof_map (SMT_Config.select_solver suffix)
|
blanchet@40941
|
541 |
#> Config.put SMT_Config.verbose debug
|
blanchet@40439
|
542 |
#> Config.put SMT_Config.monomorph_limit smt_monomorph_limit
|
blanchet@40439
|
543 |
val state = state |> Proof.map_context repair_context
|
blanchet@40668
|
544 |
val thy = Proof.theory_of state
|
blanchet@40982
|
545 |
val get_fact = Option.map (apsnd (Thm.transfer thy)) o try dest_Untranslated
|
blanchet@40983
|
546 |
val facts = facts |> map_filter get_fact
|
blanchet@40181
|
547 |
val {outcome, used_facts, run_time_in_msecs} =
|
blanchet@40982
|
548 |
smt_filter_loop params remote state subgoal facts
|
blanchet@40723
|
549 |
val (chained_lemmas, other_lemmas) = split_used_facts (map fst used_facts)
|
blanchet@40669
|
550 |
val outcome = outcome |> Option.map failure_from_smt_failure
|
blanchet@40063
|
551 |
val message =
|
blanchet@40184
|
552 |
case outcome of
|
blanchet@40184
|
553 |
NONE =>
|
blanchet@40666
|
554 |
let
|
blanchet@40666
|
555 |
val method =
|
blanchet@40693
|
556 |
if can_apply_metis debug state subgoal (map snd used_facts) then
|
blanchet@40693
|
557 |
"metis"
|
blanchet@40693
|
558 |
else
|
blanchet@40693
|
559 |
"smt"
|
blanchet@40666
|
560 |
in
|
blanchet@40666
|
561 |
try_command_line (proof_banner auto)
|
blanchet@40666
|
562 |
(apply_on_subgoal subgoal subgoal_count ^
|
blanchet@40723
|
563 |
command_call method (map fst other_lemmas)) ^
|
blanchet@40723
|
564 |
minimize_line minimize_command
|
blanchet@40723
|
565 |
(map fst (other_lemmas @ chained_lemmas))
|
blanchet@40666
|
566 |
end
|
blanchet@40669
|
567 |
| SOME failure => string_for_failure "SMT solver" failure
|
blanchet@40063
|
568 |
in
|
blanchet@40666
|
569 |
{outcome = outcome, used_facts = map fst used_facts,
|
blanchet@40204
|
570 |
run_time_in_msecs = run_time_in_msecs, message = message}
|
blanchet@40063
|
571 |
end
|
blanchet@40063
|
572 |
|
blanchet@40941
|
573 |
fun get_prover ctxt auto name =
|
blanchet@40941
|
574 |
let val thy = ProofContext.theory_of ctxt in
|
blanchet@40941
|
575 |
if is_smt_prover ctxt name then
|
blanchet@40941
|
576 |
run_smt_solver auto name
|
blanchet@40941
|
577 |
else if member (op =) (available_atps thy) name then
|
blanchet@40941
|
578 |
run_atp auto name (get_atp thy name)
|
blanchet@40941
|
579 |
else
|
blanchet@40941
|
580 |
error ("No such prover: " ^ name ^ ".")
|
blanchet@40941
|
581 |
end
|
blanchet@40063
|
582 |
|
blanchet@40063
|
583 |
fun run_prover (params as {blocking, debug, max_relevant, timeout, expect, ...})
|
blanchet@40983
|
584 |
auto minimize_command only
|
blanchet@40983
|
585 |
{state, goal, subgoal, subgoal_count, facts} name =
|
blanchet@40063
|
586 |
let
|
blanchet@39318
|
587 |
val ctxt = Proof.context_of state
|
blanchet@37584
|
588 |
val birth_time = Time.now ()
|
blanchet@37584
|
589 |
val death_time = Time.+ (birth_time, timeout)
|
blanchet@40063
|
590 |
val max_relevant =
|
blanchet@40941
|
591 |
the_default (default_max_relevant_for_prover ctxt name) max_relevant
|
blanchet@40983
|
592 |
val num_facts = length facts |> not only ? Integer.min max_relevant
|
blanchet@40065
|
593 |
val desc =
|
blanchet@40204
|
594 |
prover_description ctxt params name num_facts subgoal subgoal_count goal
|
blanchet@40941
|
595 |
val prover = get_prover ctxt auto name
|
blanchet@40983
|
596 |
val problem =
|
blanchet@40983
|
597 |
{state = state, goal = goal, subgoal = subgoal,
|
blanchet@40983
|
598 |
subgoal_count = subgoal_count, facts = take num_facts facts}
|
blanchet@39318
|
599 |
fun go () =
|
blanchet@38982
|
600 |
let
|
blanchet@39453
|
601 |
fun really_go () =
|
blanchet@40941
|
602 |
prover params (minimize_command name) problem
|
blanchet@38985
|
603 |
|> (fn {outcome, message, ...} =>
|
blanchet@38985
|
604 |
(if is_some outcome then "none" else "some", message))
|
blanchet@39453
|
605 |
val (outcome_code, message) =
|
blanchet@39453
|
606 |
if debug then
|
blanchet@39453
|
607 |
really_go ()
|
blanchet@39453
|
608 |
else
|
blanchet@39453
|
609 |
(really_go ()
|
blanchet@39453
|
610 |
handle ERROR message => ("unknown", "Error: " ^ message ^ "\n")
|
blanchet@40978
|
611 |
| exn =>
|
blanchet@40978
|
612 |
if Exn.is_interrupt exn then
|
blanchet@40978
|
613 |
reraise exn
|
blanchet@40978
|
614 |
else
|
blanchet@40978
|
615 |
("unknown", "Internal error:\n" ^
|
blanchet@40978
|
616 |
ML_Compiler.exn_message exn ^ "\n"))
|
blanchet@39318
|
617 |
val _ =
|
blanchet@39318
|
618 |
if expect = "" orelse outcome_code = expect then
|
blanchet@39318
|
619 |
()
|
blanchet@39318
|
620 |
else if blocking then
|
blanchet@39318
|
621 |
error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
|
blanchet@39318
|
622 |
else
|
blanchet@39318
|
623 |
warning ("Unexpected outcome: " ^ quote outcome_code ^ ".");
|
blanchet@39318
|
624 |
in (outcome_code = "some", message) end
|
blanchet@39318
|
625 |
in
|
blanchet@39318
|
626 |
if auto then
|
blanchet@39318
|
627 |
let val (success, message) = TimeLimit.timeLimit timeout go () in
|
blanchet@39318
|
628 |
(success, state |> success ? Proof.goal_message (fn () =>
|
blanchet@39318
|
629 |
Pretty.chunks [Pretty.str "", Pretty.mark Markup.hilite
|
blanchet@39327
|
630 |
(Pretty.str message)]))
|
blanchet@38982
|
631 |
end
|
blanchet@39318
|
632 |
else if blocking then
|
blanchet@39318
|
633 |
let val (success, message) = TimeLimit.timeLimit timeout go () in
|
wenzelm@40132
|
634 |
List.app Output.urgent_message
|
blanchet@39370
|
635 |
(Async_Manager.break_into_chunks [desc ^ "\n" ^ message]);
|
blanchet@39370
|
636 |
(success, state)
|
blanchet@39318
|
637 |
end
|
blanchet@39318
|
638 |
else
|
blanchet@39318
|
639 |
(Async_Manager.launch das_Tool birth_time death_time desc (snd o go);
|
blanchet@39318
|
640 |
(false, state))
|
blanchet@37584
|
641 |
end
|
wenzelm@28582
|
642 |
|
blanchet@38023
|
643 |
val setup =
|
blanchet@38023
|
644 |
dest_dir_setup
|
blanchet@38023
|
645 |
#> problem_prefix_setup
|
blanchet@39003
|
646 |
#> measure_run_time_setup
|
blanchet@38023
|
647 |
|
wenzelm@28582
|
648 |
end;
|