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@53586
|
11 |
type atp_failure = ATP_Proof.atp_failure
|
blanchet@46340
|
12 |
type stature = ATP_Problem_Generate.stature
|
blanchet@46320
|
13 |
type type_enc = ATP_Problem_Generate.type_enc
|
blanchet@51005
|
14 |
type fact = Sledgehammer_Fact.fact
|
smolkas@52555
|
15 |
type reconstructor = Sledgehammer_Reconstructor.reconstructor
|
smolkas@52555
|
16 |
type play = Sledgehammer_Reconstructor.play
|
smolkas@52555
|
17 |
type minimize_command = Sledgehammer_Reconstructor.minimize_command
|
blanchet@39493
|
18 |
|
wenzelm@53052
|
19 |
datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize
|
blanchet@43021
|
20 |
|
blanchet@35969
|
21 |
type params =
|
blanchet@48321
|
22 |
{debug : bool,
|
blanchet@48321
|
23 |
verbose : bool,
|
blanchet@48321
|
24 |
overlord : bool,
|
blanchet@53800
|
25 |
spy : bool,
|
blanchet@48321
|
26 |
blocking : bool,
|
blanchet@48321
|
27 |
provers : string list,
|
blanchet@48321
|
28 |
type_enc : string option,
|
blanchet@48321
|
29 |
strict : bool,
|
blanchet@48321
|
30 |
lam_trans : string option,
|
blanchet@48321
|
31 |
uncurried_aliases : bool option,
|
blanchet@48321
|
32 |
learn : bool,
|
blanchet@48321
|
33 |
fact_filter : string option,
|
blanchet@48321
|
34 |
max_facts : int option,
|
blanchet@48321
|
35 |
fact_thresholds : real * real,
|
blanchet@48321
|
36 |
max_mono_iters : int option,
|
blanchet@48321
|
37 |
max_new_mono_instances : int option,
|
blanchet@51190
|
38 |
isar_proofs : bool option,
|
smolkas@51130
|
39 |
isar_compress : real,
|
smolkas@52632
|
40 |
isar_try0 : bool,
|
blanchet@48321
|
41 |
slice : bool,
|
blanchet@48321
|
42 |
minimize : bool option,
|
blanchet@50557
|
43 |
timeout : Time.time option,
|
blanchet@50557
|
44 |
preplay_timeout : Time.time option,
|
blanchet@48321
|
45 |
expect : string}
|
blanchet@39493
|
46 |
|
blanchet@48288
|
47 |
type relevance_fudge =
|
blanchet@48288
|
48 |
{local_const_multiplier : real,
|
blanchet@48288
|
49 |
worse_irrel_freq : real,
|
blanchet@48288
|
50 |
higher_order_irrel_weight : real,
|
blanchet@48288
|
51 |
abs_rel_weight : real,
|
blanchet@48288
|
52 |
abs_irrel_weight : real,
|
blanchet@48288
|
53 |
theory_const_rel_weight : real,
|
blanchet@48288
|
54 |
theory_const_irrel_weight : real,
|
blanchet@48288
|
55 |
chained_const_irrel_weight : real,
|
blanchet@48288
|
56 |
intro_bonus : real,
|
blanchet@48288
|
57 |
elim_bonus : real,
|
blanchet@48288
|
58 |
simp_bonus : real,
|
blanchet@48288
|
59 |
local_bonus : real,
|
blanchet@48288
|
60 |
assum_bonus : real,
|
blanchet@48288
|
61 |
chained_bonus : real,
|
blanchet@48288
|
62 |
max_imperfect : real,
|
blanchet@48288
|
63 |
max_imperfect_exp : real,
|
blanchet@48288
|
64 |
threshold_divisor : real,
|
blanchet@48288
|
65 |
ridiculous_threshold : real}
|
blanchet@48288
|
66 |
|
blanchet@40061
|
67 |
type prover_problem =
|
blanchet@48321
|
68 |
{state : Proof.state,
|
blanchet@48321
|
69 |
goal : thm,
|
blanchet@48321
|
70 |
subgoal : int,
|
blanchet@48321
|
71 |
subgoal_count : int,
|
blanchet@51010
|
72 |
factss : (string * fact list) list}
|
blanchet@39493
|
73 |
|
blanchet@40061
|
74 |
type prover_result =
|
blanchet@53586
|
75 |
{outcome : atp_failure option,
|
blanchet@51009
|
76 |
used_facts : (string * stature) list,
|
blanchet@51009
|
77 |
used_from : fact list,
|
blanchet@48321
|
78 |
run_time : Time.time,
|
blanchet@50669
|
79 |
preplay : play Lazy.lazy,
|
blanchet@48321
|
80 |
message : play -> string,
|
blanchet@48321
|
81 |
message_tail : string}
|
blanchet@39493
|
82 |
|
blanchet@43051
|
83 |
type prover =
|
blanchet@45520
|
84 |
params -> ((string * string list) list -> string -> minimize_command)
|
blanchet@45520
|
85 |
-> prover_problem -> prover_result
|
blanchet@35867
|
86 |
|
blanchet@43092
|
87 |
val dest_dir : string Config.T
|
blanchet@43092
|
88 |
val problem_prefix : string Config.T
|
blanchet@48143
|
89 |
val completish : bool Config.T
|
blanchet@44592
|
90 |
val atp_full_names : bool Config.T
|
blanchet@54000
|
91 |
val smt_builtin_facts : bool Config.T
|
blanchet@54000
|
92 |
val smt_builtin_trans : bool Config.T
|
blanchet@42646
|
93 |
val smt_triggers : bool Config.T
|
blanchet@42646
|
94 |
val smt_weights : bool Config.T
|
blanchet@42646
|
95 |
val smt_weight_min_facts : int Config.T
|
blanchet@42646
|
96 |
val smt_min_weight : int Config.T
|
blanchet@42646
|
97 |
val smt_max_weight : int Config.T
|
blanchet@42646
|
98 |
val smt_max_weight_index : int Config.T
|
blanchet@41256
|
99 |
val smt_weight_curve : (int -> int) Unsynchronized.ref
|
blanchet@42646
|
100 |
val smt_max_slices : int Config.T
|
blanchet@42646
|
101 |
val smt_slice_fact_frac : real Config.T
|
blanchet@42646
|
102 |
val smt_slice_time_frac : real Config.T
|
blanchet@42646
|
103 |
val smt_slice_min_secs : int Config.T
|
blanchet@48319
|
104 |
val SledgehammerN : string
|
blanchet@45519
|
105 |
val plain_metis : reconstructor
|
blanchet@41242
|
106 |
val select_smt_solver : string -> Proof.context -> Proof.context
|
blanchet@45520
|
107 |
val extract_reconstructor :
|
blanchet@45561
|
108 |
params -> reconstructor -> string * (string * string list) list
|
blanchet@45379
|
109 |
val is_reconstructor : string -> bool
|
blanchet@43050
|
110 |
val is_atp : theory -> string -> bool
|
blanchet@41087
|
111 |
val is_smt_prover : Proof.context -> string -> bool
|
blanchet@47962
|
112 |
val is_ho_atp: Proof.context -> string -> bool
|
blanchet@42944
|
113 |
val is_unit_equational_atp : Proof.context -> string -> bool
|
blanchet@41727
|
114 |
val is_prover_supported : Proof.context -> string -> bool
|
blanchet@40072
|
115 |
val is_prover_installed : Proof.context -> string -> bool
|
blanchet@51200
|
116 |
val remotify_prover_if_supported_and_not_already_remote :
|
blanchet@51200
|
117 |
Proof.context -> string -> string option
|
blanchet@51200
|
118 |
val remotify_prover_if_not_installed :
|
blanchet@51200
|
119 |
Proof.context -> string -> string option
|
blanchet@51998
|
120 |
val default_max_facts_of_prover : Proof.context -> bool -> string -> int
|
blanchet@42944
|
121 |
val is_unit_equality : term -> bool
|
blanchet@51998
|
122 |
val is_appropriate_prop_of_prover : Proof.context -> string -> term -> bool
|
blanchet@51998
|
123 |
val is_built_in_const_of_prover :
|
blanchet@41336
|
124 |
Proof.context -> string -> string * typ -> term list -> bool * term list
|
blanchet@41087
|
125 |
val atp_relevance_fudge : relevance_fudge
|
blanchet@41087
|
126 |
val smt_relevance_fudge : relevance_fudge
|
blanchet@51998
|
127 |
val relevance_fudge_of_prover : Proof.context -> string -> relevance_fudge
|
blanchet@41256
|
128 |
val weight_smt_fact :
|
blanchet@46340
|
129 |
Proof.context -> int -> ((string * stature) * thm) * int
|
blanchet@46340
|
130 |
-> (string * stature) * (int option * thm)
|
blanchet@41727
|
131 |
val supported_provers : Proof.context -> unit
|
blanchet@40059
|
132 |
val kill_provers : unit -> unit
|
blanchet@40059
|
133 |
val running_provers : unit -> unit
|
blanchet@40059
|
134 |
val messages : int option -> unit
|
blanchet@48798
|
135 |
val is_fact_chained : (('a * stature) * 'b) -> bool
|
blanchet@48798
|
136 |
val filter_used_facts :
|
blanchet@48798
|
137 |
bool -> (''a * stature) list -> ((''a * stature) * 'b) list ->
|
blanchet@48798
|
138 |
((''a * stature) * 'b) list
|
blanchet@51200
|
139 |
val isar_proof_reconstructor : Proof.context -> string -> string
|
blanchet@43021
|
140 |
val get_prover : Proof.context -> mode -> string -> prover
|
wenzelm@28477
|
141 |
end;
|
wenzelm@28477
|
142 |
|
blanchet@41087
|
143 |
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS =
|
wenzelm@28477
|
144 |
struct
|
wenzelm@28477
|
145 |
|
blanchet@43085
|
146 |
open ATP_Util
|
blanchet@38028
|
147 |
open ATP_Problem
|
blanchet@39491
|
148 |
open ATP_Proof
|
blanchet@38028
|
149 |
open ATP_Systems
|
blanchet@46320
|
150 |
open ATP_Problem_Generate
|
blanchet@46320
|
151 |
open ATP_Proof_Reconstruct
|
blanchet@45521
|
152 |
open Metis_Tactic
|
blanchet@38023
|
153 |
open Sledgehammer_Util
|
blanchet@51005
|
154 |
open Sledgehammer_Fact
|
smolkas@52555
|
155 |
open Sledgehammer_Reconstructor
|
smolkas@52555
|
156 |
open Sledgehammer_Print
|
blanchet@49881
|
157 |
open Sledgehammer_Reconstruct
|
blanchet@48288
|
158 |
|
blanchet@37583
|
159 |
|
blanchet@37583
|
160 |
(** The Sledgehammer **)
|
blanchet@37583
|
161 |
|
blanchet@54000
|
162 |
(* Empty string means create files in Isabelle's temporary files directory. *)
|
blanchet@54000
|
163 |
val dest_dir =
|
blanchet@54000
|
164 |
Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
|
blanchet@54000
|
165 |
val problem_prefix =
|
blanchet@54000
|
166 |
Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
|
blanchet@54000
|
167 |
val completish =
|
blanchet@54000
|
168 |
Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
|
blanchet@54000
|
169 |
|
blanchet@54000
|
170 |
(* In addition to being easier to read, readable names are often much shorter,
|
blanchet@54000
|
171 |
especially if types are mangled in names. This makes a difference for some
|
blanchet@54000
|
172 |
provers (e.g., E). For these reason, short names are enabled by default. *)
|
blanchet@54000
|
173 |
val atp_full_names =
|
blanchet@54000
|
174 |
Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
|
blanchet@54000
|
175 |
|
blanchet@54000
|
176 |
val smt_builtin_facts =
|
blanchet@54000
|
177 |
Attrib.setup_config_bool @{binding sledgehammer_smt_builtin_facts} (K true)
|
blanchet@54000
|
178 |
val smt_builtin_trans =
|
blanchet@54000
|
179 |
Attrib.setup_config_bool @{binding sledgehammer_smt_builtin_trans} (K true)
|
blanchet@54000
|
180 |
val smt_triggers =
|
blanchet@54000
|
181 |
Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
|
blanchet@54000
|
182 |
val smt_weights =
|
blanchet@54000
|
183 |
Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
|
blanchet@54000
|
184 |
val smt_weight_min_facts =
|
blanchet@54000
|
185 |
Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
|
blanchet@54000
|
186 |
|
wenzelm@53052
|
187 |
datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize
|
blanchet@43021
|
188 |
|
blanchet@45376
|
189 |
(* Identifier that distinguishes Sledgehammer from other tools that could use
|
blanchet@38102
|
190 |
"Async_Manager". *)
|
blanchet@48319
|
191 |
val SledgehammerN = "Sledgehammer"
|
blanchet@37585
|
192 |
|
blanchet@45520
|
193 |
val reconstructor_names = [metisN, smtN]
|
blanchet@46365
|
194 |
val plain_metis = Metis (hd partial_type_encs, combsN)
|
blanchet@45561
|
195 |
val is_reconstructor = member (op =) reconstructor_names
|
blanchet@43228
|
196 |
|
blanchet@43050
|
197 |
val is_atp = member (op =) o supported_atps
|
blanchet@43050
|
198 |
|
blanchet@43233
|
199 |
val select_smt_solver = Context.proof_map o SMT_Config.select_solver
|
blanchet@41242
|
200 |
|
blanchet@45376
|
201 |
fun is_smt_prover ctxt = member (op =) (SMT_Solver.available_solvers_of ctxt)
|
blanchet@40062
|
202 |
|
blanchet@51998
|
203 |
fun is_atp_of_format is_format ctxt name =
|
blanchet@42944
|
204 |
let val thy = Proof_Context.theory_of ctxt in
|
blanchet@42944
|
205 |
case try (get_atp thy) name of
|
blanchet@47606
|
206 |
SOME config =>
|
blanchet@48716
|
207 |
exists (fn (_, ((_, format, _, _, _), _)) => is_format format)
|
blanchet@47606
|
208 |
(#best_slices (config ()) ctxt)
|
blanchet@42944
|
209 |
| NONE => false
|
blanchet@42944
|
210 |
end
|
blanchet@42944
|
211 |
|
blanchet@51998
|
212 |
val is_unit_equational_atp = is_atp_of_format (curry (op =) CNF_UEQ)
|
blanchet@51998
|
213 |
val is_ho_atp = is_atp_of_format is_format_higher_order
|
blanchet@44597
|
214 |
|
blanchet@45376
|
215 |
fun is_prover_supported ctxt =
|
wenzelm@42361
|
216 |
let val thy = Proof_Context.theory_of ctxt in
|
blanchet@45379
|
217 |
is_reconstructor orf is_atp thy orf is_smt_prover ctxt
|
blanchet@40941
|
218 |
end
|
blanchet@40072
|
219 |
|
boehmes@41432
|
220 |
fun is_prover_installed ctxt =
|
blanchet@45379
|
221 |
is_reconstructor orf is_smt_prover ctxt orf
|
blanchet@43050
|
222 |
is_atp_installed (Proof_Context.theory_of ctxt)
|
blanchet@40941
|
223 |
|
blanchet@51200
|
224 |
fun remotify_prover_if_supported_and_not_already_remote ctxt name =
|
blanchet@51200
|
225 |
if String.isPrefix remote_prefix name then
|
blanchet@51200
|
226 |
SOME name
|
blanchet@51200
|
227 |
else
|
blanchet@51200
|
228 |
let val remote_name = remote_prefix ^ name in
|
blanchet@51200
|
229 |
if is_prover_supported ctxt remote_name then SOME remote_name else NONE
|
blanchet@51200
|
230 |
end
|
blanchet@51200
|
231 |
|
blanchet@51200
|
232 |
fun remotify_prover_if_not_installed ctxt name =
|
blanchet@51200
|
233 |
if is_prover_supported ctxt name andalso is_prover_installed ctxt name then
|
blanchet@51200
|
234 |
SOME name
|
blanchet@51200
|
235 |
else
|
blanchet@51200
|
236 |
remotify_prover_if_supported_and_not_already_remote ctxt name
|
blanchet@51200
|
237 |
|
blanchet@45706
|
238 |
fun get_slices slice slices =
|
blanchet@45706
|
239 |
(0 upto length slices - 1) ~~ slices |> not slice ? (List.last #> single)
|
blanchet@42443
|
240 |
|
blanchet@48293
|
241 |
val reconstructor_default_max_facts = 20
|
blanchet@43050
|
242 |
|
blanchet@51186
|
243 |
fun slice_max_facts (_, (_, ( ((max_facts, _), _, _, _, _), _))) = max_facts
|
blanchet@51186
|
244 |
|
blanchet@51998
|
245 |
fun default_max_facts_of_prover ctxt slice name =
|
wenzelm@42361
|
246 |
let val thy = Proof_Context.theory_of ctxt in
|
blanchet@45379
|
247 |
if is_reconstructor name then
|
blanchet@48293
|
248 |
reconstructor_default_max_facts
|
blanchet@43050
|
249 |
else if is_atp thy name then
|
blanchet@51186
|
250 |
fold (Integer.max o slice_max_facts)
|
blanchet@47606
|
251 |
(get_slices slice (#best_slices (get_atp thy name ()) ctxt)) 0
|
blanchet@43050
|
252 |
else (* is_smt_prover ctxt name *)
|
blanchet@43050
|
253 |
SMT_Solver.default_max_relevant ctxt name
|
blanchet@40941
|
254 |
end
|
blanchet@40063
|
255 |
|
blanchet@42956
|
256 |
fun is_if (@{const_name If}, _) = true
|
blanchet@42956
|
257 |
| is_if _ = false
|
blanchet@42956
|
258 |
|
blanchet@42956
|
259 |
(* Beware of "if and only if" (which is translated as such) and "If" (which is
|
blanchet@42956
|
260 |
translated to conditional equations). *)
|
blanchet@42956
|
261 |
fun is_good_unit_equality T t u =
|
blanchet@42956
|
262 |
T <> @{typ bool} andalso not (exists (exists_Const is_if) [t, u])
|
blanchet@42956
|
263 |
|
blanchet@42944
|
264 |
fun is_unit_equality (@{const Trueprop} $ t) = is_unit_equality t
|
blanchet@42944
|
265 |
| is_unit_equality (Const (@{const_name all}, _) $ Abs (_, _, t)) =
|
blanchet@42944
|
266 |
is_unit_equality t
|
blanchet@42944
|
267 |
| is_unit_equality (Const (@{const_name All}, _) $ Abs (_, _, t)) =
|
blanchet@42944
|
268 |
is_unit_equality t
|
blanchet@42956
|
269 |
| is_unit_equality (Const (@{const_name "=="}, Type (_, [T, _])) $ t $ u) =
|
blanchet@42956
|
270 |
is_good_unit_equality T t u
|
blanchet@42956
|
271 |
| is_unit_equality (Const (@{const_name HOL.eq}, Type (_ , [T, _])) $ t $ u) =
|
blanchet@42956
|
272 |
is_good_unit_equality T t u
|
blanchet@42944
|
273 |
| is_unit_equality _ = false
|
blanchet@42944
|
274 |
|
blanchet@51998
|
275 |
fun is_appropriate_prop_of_prover ctxt name =
|
blanchet@42944
|
276 |
if is_unit_equational_atp ctxt name then is_unit_equality else K true
|
blanchet@42944
|
277 |
|
blanchet@53517
|
278 |
val atp_irrelevant_const_tab =
|
blanchet@53517
|
279 |
Symtab.make (map (rpair ()) atp_irrelevant_consts)
|
blanchet@53517
|
280 |
|
blanchet@51998
|
281 |
fun is_built_in_const_of_prover ctxt name =
|
blanchet@54000
|
282 |
if is_smt_prover ctxt name andalso Config.get ctxt smt_builtin_facts then
|
blanchet@41336
|
283 |
let val ctxt = ctxt |> select_smt_solver name in
|
blanchet@41336
|
284 |
fn x => fn ts =>
|
blanchet@41336
|
285 |
if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then
|
blanchet@41336
|
286 |
(true, [])
|
blanchet@41336
|
287 |
else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then
|
blanchet@41336
|
288 |
(true, ts)
|
blanchet@41336
|
289 |
else
|
blanchet@41336
|
290 |
(false, ts)
|
blanchet@41336
|
291 |
end
|
blanchet@41242
|
292 |
else
|
blanchet@53517
|
293 |
fn (s, _) => fn ts => (Symtab.defined atp_irrelevant_const_tab s, ts)
|
blanchet@40071
|
294 |
|
blanchet@40070
|
295 |
(* FUDGE *)
|
blanchet@40070
|
296 |
val atp_relevance_fudge =
|
blanchet@42738
|
297 |
{local_const_multiplier = 1.5,
|
blanchet@41159
|
298 |
worse_irrel_freq = 100.0,
|
blanchet@40070
|
299 |
higher_order_irrel_weight = 1.05,
|
blanchet@40070
|
300 |
abs_rel_weight = 0.5,
|
blanchet@40070
|
301 |
abs_irrel_weight = 2.0,
|
blanchet@40070
|
302 |
theory_const_rel_weight = 0.5,
|
blanchet@40070
|
303 |
theory_const_irrel_weight = 0.25,
|
blanchet@42735
|
304 |
chained_const_irrel_weight = 0.25,
|
blanchet@40070
|
305 |
intro_bonus = 0.15,
|
blanchet@40070
|
306 |
elim_bonus = 0.15,
|
blanchet@40070
|
307 |
simp_bonus = 0.15,
|
blanchet@40070
|
308 |
local_bonus = 0.55,
|
blanchet@40070
|
309 |
assum_bonus = 1.05,
|
blanchet@40070
|
310 |
chained_bonus = 1.5,
|
blanchet@40070
|
311 |
max_imperfect = 11.5,
|
blanchet@40070
|
312 |
max_imperfect_exp = 1.0,
|
blanchet@40070
|
313 |
threshold_divisor = 2.0,
|
blanchet@41093
|
314 |
ridiculous_threshold = 0.01}
|
blanchet@40070
|
315 |
|
blanchet@40071
|
316 |
(* FUDGE (FIXME) *)
|
blanchet@40070
|
317 |
val smt_relevance_fudge =
|
blanchet@42738
|
318 |
{local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
|
blanchet@41159
|
319 |
worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
|
blanchet@40071
|
320 |
higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
|
blanchet@40071
|
321 |
abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
|
blanchet@40071
|
322 |
abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
|
blanchet@40071
|
323 |
theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
|
blanchet@40071
|
324 |
theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
|
blanchet@42735
|
325 |
chained_const_irrel_weight = #chained_const_irrel_weight atp_relevance_fudge,
|
blanchet@40071
|
326 |
intro_bonus = #intro_bonus atp_relevance_fudge,
|
blanchet@40071
|
327 |
elim_bonus = #elim_bonus atp_relevance_fudge,
|
blanchet@40071
|
328 |
simp_bonus = #simp_bonus atp_relevance_fudge,
|
blanchet@40071
|
329 |
local_bonus = #local_bonus atp_relevance_fudge,
|
blanchet@40071
|
330 |
assum_bonus = #assum_bonus atp_relevance_fudge,
|
blanchet@40071
|
331 |
chained_bonus = #chained_bonus atp_relevance_fudge,
|
blanchet@40071
|
332 |
max_imperfect = #max_imperfect atp_relevance_fudge,
|
blanchet@40071
|
333 |
max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
|
blanchet@40071
|
334 |
threshold_divisor = #threshold_divisor atp_relevance_fudge,
|
blanchet@40071
|
335 |
ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
|
blanchet@40070
|
336 |
|
blanchet@51998
|
337 |
fun relevance_fudge_of_prover ctxt name =
|
blanchet@40941
|
338 |
if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
|
blanchet@40070
|
339 |
|
blanchet@41727
|
340 |
fun supported_provers ctxt =
|
blanchet@40060
|
341 |
let
|
wenzelm@42361
|
342 |
val thy = Proof_Context.theory_of ctxt
|
blanchet@40060
|
343 |
val (remote_provers, local_provers) =
|
blanchet@45520
|
344 |
reconstructor_names @
|
blanchet@41727
|
345 |
sort_strings (supported_atps thy) @
|
blanchet@41727
|
346 |
sort_strings (SMT_Solver.available_solvers_of ctxt)
|
blanchet@40060
|
347 |
|> List.partition (String.isPrefix remote_prefix)
|
blanchet@40060
|
348 |
in
|
blanchet@41727
|
349 |
Output.urgent_message ("Supported provers: " ^
|
blanchet@40205
|
350 |
commas (local_provers @ remote_provers) ^ ".")
|
blanchet@40060
|
351 |
end
|
blanchet@35969
|
352 |
|
blanchet@48319
|
353 |
fun kill_provers () = Async_Manager.kill_threads SledgehammerN "prover"
|
blanchet@48319
|
354 |
fun running_provers () = Async_Manager.running_threads SledgehammerN "prover"
|
blanchet@48319
|
355 |
val messages = Async_Manager.thread_messages SledgehammerN "prover"
|
blanchet@40059
|
356 |
|
blanchet@48288
|
357 |
|
blanchet@40059
|
358 |
(** problems, results, ATPs, etc. **)
|
blanchet@35969
|
359 |
|
blanchet@35969
|
360 |
type params =
|
blanchet@48321
|
361 |
{debug : bool,
|
blanchet@48321
|
362 |
verbose : bool,
|
blanchet@48321
|
363 |
overlord : bool,
|
blanchet@53800
|
364 |
spy : bool,
|
blanchet@48321
|
365 |
blocking : bool,
|
blanchet@48321
|
366 |
provers : string list,
|
blanchet@48321
|
367 |
type_enc : string option,
|
blanchet@48321
|
368 |
strict : bool,
|
blanchet@48321
|
369 |
lam_trans : string option,
|
blanchet@48321
|
370 |
uncurried_aliases : bool option,
|
blanchet@48321
|
371 |
learn : bool,
|
blanchet@48321
|
372 |
fact_filter : string option,
|
blanchet@48321
|
373 |
max_facts : int option,
|
blanchet@48321
|
374 |
fact_thresholds : real * real,
|
blanchet@48321
|
375 |
max_mono_iters : int option,
|
blanchet@48321
|
376 |
max_new_mono_instances : int option,
|
blanchet@51190
|
377 |
isar_proofs : bool option,
|
smolkas@51130
|
378 |
isar_compress : real,
|
smolkas@52632
|
379 |
isar_try0 : bool,
|
blanchet@48321
|
380 |
slice : bool,
|
blanchet@48321
|
381 |
minimize : bool option,
|
blanchet@50557
|
382 |
timeout : Time.time option,
|
blanchet@50557
|
383 |
preplay_timeout : Time.time option,
|
blanchet@48321
|
384 |
expect : string}
|
blanchet@35867
|
385 |
|
blanchet@48288
|
386 |
type relevance_fudge =
|
blanchet@48288
|
387 |
{local_const_multiplier : real,
|
blanchet@48288
|
388 |
worse_irrel_freq : real,
|
blanchet@48288
|
389 |
higher_order_irrel_weight : real,
|
blanchet@48288
|
390 |
abs_rel_weight : real,
|
blanchet@48288
|
391 |
abs_irrel_weight : real,
|
blanchet@48288
|
392 |
theory_const_rel_weight : real,
|
blanchet@48288
|
393 |
theory_const_irrel_weight : real,
|
blanchet@48288
|
394 |
chained_const_irrel_weight : real,
|
blanchet@48288
|
395 |
intro_bonus : real,
|
blanchet@48288
|
396 |
elim_bonus : real,
|
blanchet@48288
|
397 |
simp_bonus : real,
|
blanchet@48288
|
398 |
local_bonus : real,
|
blanchet@48288
|
399 |
assum_bonus : real,
|
blanchet@48288
|
400 |
chained_bonus : real,
|
blanchet@48288
|
401 |
max_imperfect : real,
|
blanchet@48288
|
402 |
max_imperfect_exp : real,
|
blanchet@48288
|
403 |
threshold_divisor : real,
|
blanchet@48288
|
404 |
ridiculous_threshold : real}
|
blanchet@48288
|
405 |
|
blanchet@40061
|
406 |
type prover_problem =
|
blanchet@48321
|
407 |
{state : Proof.state,
|
blanchet@48321
|
408 |
goal : thm,
|
blanchet@48321
|
409 |
subgoal : int,
|
blanchet@48321
|
410 |
subgoal_count : int,
|
blanchet@51010
|
411 |
factss : (string * fact list) list}
|
blanchet@35867
|
412 |
|
blanchet@40061
|
413 |
type prover_result =
|
blanchet@53586
|
414 |
{outcome : atp_failure option,
|
blanchet@48321
|
415 |
used_facts : (string * stature) list,
|
blanchet@51009
|
416 |
used_from : fact list,
|
blanchet@48321
|
417 |
run_time : Time.time,
|
blanchet@50669
|
418 |
preplay : play Lazy.lazy,
|
blanchet@48321
|
419 |
message : play -> string,
|
blanchet@48321
|
420 |
message_tail : string}
|
blanchet@35867
|
421 |
|
blanchet@43051
|
422 |
type prover =
|
blanchet@45520
|
423 |
params -> ((string * string list) list -> string -> minimize_command)
|
blanchet@45520
|
424 |
-> prover_problem -> prover_result
|
blanchet@35867
|
425 |
|
blanchet@41256
|
426 |
(* FUDGE *)
|
blanchet@42646
|
427 |
val smt_min_weight =
|
blanchet@42646
|
428 |
Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
|
blanchet@42646
|
429 |
val smt_max_weight =
|
blanchet@42646
|
430 |
Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
|
blanchet@42646
|
431 |
val smt_max_weight_index =
|
blanchet@42646
|
432 |
Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
|
blanchet@41256
|
433 |
val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x)
|
blanchet@41256
|
434 |
|
blanchet@42646
|
435 |
fun smt_fact_weight ctxt j num_facts =
|
blanchet@42646
|
436 |
if Config.get ctxt smt_weights andalso
|
blanchet@42646
|
437 |
num_facts >= Config.get ctxt smt_weight_min_facts then
|
blanchet@42646
|
438 |
let
|
blanchet@42646
|
439 |
val min = Config.get ctxt smt_min_weight
|
blanchet@42646
|
440 |
val max = Config.get ctxt smt_max_weight
|
blanchet@42646
|
441 |
val max_index = Config.get ctxt smt_max_weight_index
|
blanchet@42646
|
442 |
val curve = !smt_weight_curve
|
blanchet@42646
|
443 |
in
|
blanchet@42646
|
444 |
SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1))
|
blanchet@42646
|
445 |
div curve max_index)
|
blanchet@42646
|
446 |
end
|
blanchet@41256
|
447 |
else
|
blanchet@41256
|
448 |
NONE
|
blanchet@41256
|
449 |
|
blanchet@42646
|
450 |
fun weight_smt_fact ctxt num_facts ((info, th), j) =
|
blanchet@42646
|
451 |
let val thy = Proof_Context.theory_of ctxt in
|
blanchet@42646
|
452 |
(info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy))
|
blanchet@42646
|
453 |
end
|
blanchet@38023
|
454 |
|
blanchet@51998
|
455 |
fun overlord_file_location_of_prover prover =
|
blanchet@41313
|
456 |
(getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
|
blanchet@41313
|
457 |
|
blanchet@43052
|
458 |
fun proof_banner mode name =
|
blanchet@43033
|
459 |
case mode of
|
blanchet@43033
|
460 |
Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof"
|
blanchet@43033
|
461 |
| Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
|
blanchet@43052
|
462 |
| _ => "Try this"
|
blanchet@43033
|
463 |
|
blanchet@45561
|
464 |
fun bunch_of_reconstructors needs_full_types lam_trans =
|
blanchet@48800
|
465 |
if needs_full_types then
|
blanchet@48802
|
466 |
[Metis (full_type_enc, lam_trans false),
|
blanchet@48802
|
467 |
Metis (really_full_type_enc, lam_trans false),
|
blanchet@48802
|
468 |
Metis (full_type_enc, lam_trans true),
|
blanchet@48802
|
469 |
Metis (really_full_type_enc, lam_trans true),
|
blanchet@48802
|
470 |
SMT]
|
blanchet@48802
|
471 |
else
|
blanchet@48800
|
472 |
[Metis (partial_type_enc, lam_trans false),
|
blanchet@48800
|
473 |
Metis (full_type_enc, lam_trans false),
|
blanchet@48800
|
474 |
Metis (no_typesN, lam_trans true),
|
blanchet@48800
|
475 |
Metis (really_full_type_enc, lam_trans true),
|
blanchet@48800
|
476 |
SMT]
|
blanchet@45561
|
477 |
|
blanchet@45561
|
478 |
fun extract_reconstructor ({type_enc, lam_trans, ...} : params)
|
blanchet@45561
|
479 |
(Metis (type_enc', lam_trans')) =
|
blanchet@45561
|
480 |
let
|
blanchet@45561
|
481 |
val override_params =
|
blanchet@45561
|
482 |
(if is_none type_enc andalso type_enc' = hd partial_type_encs then
|
blanchet@45561
|
483 |
[]
|
blanchet@45561
|
484 |
else
|
blanchet@45566
|
485 |
[("type_enc", [hd (unalias_type_enc type_enc')])]) @
|
blanchet@45561
|
486 |
(if is_none lam_trans andalso lam_trans' = metis_default_lam_trans then
|
blanchet@45561
|
487 |
[]
|
blanchet@45561
|
488 |
else
|
blanchet@45561
|
489 |
[("lam_trans", [lam_trans'])])
|
blanchet@45561
|
490 |
in (metisN, override_params) end
|
blanchet@45561
|
491 |
| extract_reconstructor _ SMT = (smtN, [])
|
blanchet@45561
|
492 |
|
blanchet@43033
|
493 |
(* based on "Mirabelle.can_apply" and generalized *)
|
blanchet@43034
|
494 |
fun timed_apply timeout tac state i =
|
blanchet@43033
|
495 |
let
|
blanchet@43033
|
496 |
val {context = ctxt, facts, goal} = Proof.goal state
|
blanchet@43033
|
497 |
val full_tac = Method.insert_tac facts i THEN tac ctxt i
|
blanchet@50557
|
498 |
in time_limit timeout (try (Seq.pull o full_tac)) goal end
|
blanchet@43033
|
499 |
|
blanchet@51998
|
500 |
fun tac_of_reconstructor (Metis (type_enc, lam_trans)) =
|
blanchet@45521
|
501 |
metis_tac [type_enc] lam_trans
|
blanchet@51998
|
502 |
| tac_of_reconstructor SMT = SMT_Solver.smt_tac
|
blanchet@43034
|
503 |
|
blanchet@45520
|
504 |
fun timed_reconstructor reconstr debug timeout ths =
|
blanchet@44651
|
505 |
(Config.put Metis_Tactic.verbose debug
|
blanchet@45557
|
506 |
#> Config.put SMT_Config.verbose debug
|
blanchet@51998
|
507 |
#> (fn ctxt => tac_of_reconstructor reconstr ctxt ths))
|
blanchet@43034
|
508 |
|> timed_apply timeout
|
blanchet@43033
|
509 |
|
blanchet@48798
|
510 |
fun is_fact_chained ((_, (sc, _)), _) = sc = Chained
|
blanchet@48798
|
511 |
|
blanchet@48798
|
512 |
fun filter_used_facts keep_chained used =
|
blanchet@48798
|
513 |
filter ((member (op =) used o fst) orf
|
blanchet@48798
|
514 |
(if keep_chained then is_fact_chained else K false))
|
blanchet@43033
|
515 |
|
blanchet@45781
|
516 |
fun play_one_line_proof mode debug verbose timeout pairs state i preferred
|
blanchet@45520
|
517 |
reconstrs =
|
blanchet@43034
|
518 |
let
|
blanchet@45520
|
519 |
fun get_preferred reconstrs =
|
blanchet@45520
|
520 |
if member (op =) reconstrs preferred then preferred
|
blanchet@45520
|
521 |
else List.last reconstrs
|
blanchet@43034
|
522 |
in
|
blanchet@50557
|
523 |
if timeout = SOME Time.zeroTime then
|
blanchet@45520
|
524 |
Trust_Playable (get_preferred reconstrs, NONE)
|
blanchet@45379
|
525 |
else
|
blanchet@50557
|
526 |
let
|
blanchet@50557
|
527 |
val _ =
|
blanchet@50557
|
528 |
if mode = Minimize then Output.urgent_message "Preplaying proof..."
|
blanchet@50557
|
529 |
else ()
|
blanchet@50557
|
530 |
val ths = pairs |> sort_wrt (fst o fst) |> map snd
|
blanchet@50557
|
531 |
fun play [] [] = Failed_to_Play (get_preferred reconstrs)
|
blanchet@50557
|
532 |
| play timed_outs [] =
|
blanchet@50557
|
533 |
Trust_Playable (get_preferred timed_outs, timeout)
|
blanchet@50557
|
534 |
| play timed_out (reconstr :: reconstrs) =
|
blanchet@50557
|
535 |
let
|
blanchet@50557
|
536 |
val _ =
|
blanchet@50557
|
537 |
if verbose then
|
blanchet@51998
|
538 |
"Trying \"" ^ string_of_reconstructor reconstr ^ "\"" ^
|
blanchet@50557
|
539 |
(case timeout of
|
blanchet@52031
|
540 |
SOME timeout => " for " ^ string_of_time timeout
|
blanchet@50557
|
541 |
| NONE => "") ^ "..."
|
blanchet@50557
|
542 |
|> Output.urgent_message
|
blanchet@50557
|
543 |
else
|
blanchet@50557
|
544 |
()
|
blanchet@50557
|
545 |
val timer = Timer.startRealTimer ()
|
blanchet@50557
|
546 |
in
|
blanchet@50557
|
547 |
case timed_reconstructor reconstr debug timeout ths state i of
|
blanchet@50557
|
548 |
SOME (SOME _) => Played (reconstr, Timer.checkRealTimer timer)
|
blanchet@50557
|
549 |
| _ => play timed_out reconstrs
|
blanchet@50557
|
550 |
end
|
blanchet@50557
|
551 |
handle TimeLimit.TimeOut => play (reconstr :: timed_out) reconstrs
|
blanchet@50557
|
552 |
in play [] reconstrs end
|
blanchet@43033
|
553 |
end
|
blanchet@43033
|
554 |
|
blanchet@41313
|
555 |
|
blanchet@41256
|
556 |
(* generic TPTP-based ATPs *)
|
blanchet@40061
|
557 |
|
blanchet@42730
|
558 |
(* Too general means, positive equality literal with a variable X as one
|
blanchet@42730
|
559 |
operand, when X does not occur properly in the other operand. This rules out
|
blanchet@42730
|
560 |
clearly inconsistent facts such as X = a | X = b, though it by no means
|
blanchet@42730
|
561 |
guarantees soundness. *)
|
blanchet@42730
|
562 |
|
blanchet@51998
|
563 |
fun get_facts_of_filter _ [(_, facts)] = facts
|
blanchet@51998
|
564 |
| get_facts_of_filter fact_filter factss =
|
blanchet@51013
|
565 |
case AList.lookup (op =) factss fact_filter of
|
blanchet@51013
|
566 |
SOME facts => facts
|
blanchet@51013
|
567 |
| NONE => snd (hd factss)
|
blanchet@51013
|
568 |
|
blanchet@42730
|
569 |
(* Unwanted equalities are those between a (bound or schematic) variable that
|
blanchet@42730
|
570 |
does not properly occur in the second operand. *)
|
blanchet@42730
|
571 |
val is_exhaustive_finite =
|
blanchet@42730
|
572 |
let
|
blanchet@42730
|
573 |
fun is_bad_equal (Var z) t =
|
blanchet@42730
|
574 |
not (exists_subterm (fn Var z' => z = z' | _ => false) t)
|
blanchet@42730
|
575 |
| is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
|
blanchet@42730
|
576 |
| is_bad_equal _ _ = false
|
blanchet@42730
|
577 |
fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
|
blanchet@42730
|
578 |
fun do_formula pos t =
|
blanchet@42730
|
579 |
case (pos, t) of
|
blanchet@42730
|
580 |
(_, @{const Trueprop} $ t1) => do_formula pos t1
|
blanchet@42730
|
581 |
| (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
|
blanchet@42730
|
582 |
do_formula pos t'
|
blanchet@42730
|
583 |
| (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
|
blanchet@42730
|
584 |
do_formula pos t'
|
blanchet@42730
|
585 |
| (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
|
blanchet@42730
|
586 |
do_formula pos t'
|
blanchet@42730
|
587 |
| (_, @{const "==>"} $ t1 $ t2) =>
|
blanchet@42730
|
588 |
do_formula (not pos) t1 andalso
|
blanchet@42730
|
589 |
(t2 = @{prop False} orelse do_formula pos t2)
|
blanchet@42730
|
590 |
| (_, @{const HOL.implies} $ t1 $ t2) =>
|
blanchet@42730
|
591 |
do_formula (not pos) t1 andalso
|
blanchet@42730
|
592 |
(t2 = @{const False} orelse do_formula pos t2)
|
blanchet@42730
|
593 |
| (_, @{const Not} $ t1) => do_formula (not pos) t1
|
blanchet@42730
|
594 |
| (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
|
blanchet@42730
|
595 |
| (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
|
blanchet@42730
|
596 |
| (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
|
blanchet@42730
|
597 |
| (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
|
blanchet@42730
|
598 |
| _ => false
|
blanchet@42730
|
599 |
in do_formula true end
|
blanchet@42730
|
600 |
|
blanchet@42730
|
601 |
fun has_bound_or_var_of_type pred =
|
blanchet@42730
|
602 |
exists_subterm (fn Var (_, T as Type _) => pred T
|
blanchet@42730
|
603 |
| Abs (_, T as Type _, _) => pred T
|
blanchet@42730
|
604 |
| _ => false)
|
blanchet@42730
|
605 |
|
blanchet@42730
|
606 |
(* Facts are forbidden to contain variables of these types. The typical reason
|
blanchet@42730
|
607 |
is that they lead to unsoundness. Note that "unit" satisfies numerous
|
blanchet@42730
|
608 |
equations like "?x = ()". The resulting clauses will have no type constraint,
|
blanchet@42730
|
609 |
yielding false proofs. Even "bool" leads to many unsound proofs, though only
|
blanchet@42730
|
610 |
for higher-order problems. *)
|
blanchet@42730
|
611 |
|
blanchet@42730
|
612 |
(* Facts containing variables of type "unit" or "bool" or of the form
|
blanchet@42730
|
613 |
"ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
|
blanchet@42730
|
614 |
are omitted. *)
|
blanchet@42944
|
615 |
fun is_dangerous_prop ctxt =
|
blanchet@42944
|
616 |
transform_elim_prop
|
blanchet@44393
|
617 |
#> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf
|
blanchet@42730
|
618 |
is_exhaustive_finite)
|
blanchet@42730
|
619 |
|
blanchet@39492
|
620 |
(* Important messages are important but not so important that users want to see
|
blanchet@39492
|
621 |
them each time. *)
|
blanchet@44649
|
622 |
val atp_important_message_keep_quotient = 25
|
blanchet@39492
|
623 |
|
blanchet@44416
|
624 |
fun choose_type_enc soundness best_type_enc format =
|
blanchet@44397
|
625 |
the_default best_type_enc
|
blanchet@52031
|
626 |
#> type_enc_of_string soundness
|
blanchet@44416
|
627 |
#> adjust_type_enc format
|
blanchet@42548
|
628 |
|
blanchet@51200
|
629 |
fun isar_proof_reconstructor ctxt name =
|
blanchet@51200
|
630 |
let val thy = Proof_Context.theory_of ctxt in
|
blanchet@51200
|
631 |
if is_atp thy name then name
|
blanchet@51200
|
632 |
else remotify_prover_if_not_installed ctxt eN |> the_default name
|
blanchet@51200
|
633 |
end
|
blanchet@43051
|
634 |
|
blanchet@51200
|
635 |
(* See the analogous logic in the function "maybe_minimize" in
|
blanchet@51200
|
636 |
"sledgehammer_minimize.ML". *)
|
blanchet@51200
|
637 |
fun choose_minimize_command ctxt (params as {isar_proofs, ...}) minimize_command
|
blanchet@51200
|
638 |
name preplay =
|
blanchet@45520
|
639 |
let
|
blanchet@51200
|
640 |
val maybe_isar_name =
|
blanchet@51200
|
641 |
name |> isar_proofs = SOME true ? isar_proof_reconstructor ctxt
|
blanchet@51200
|
642 |
val (min_name, override_params) =
|
blanchet@45520
|
643 |
case preplay of
|
blanchet@51200
|
644 |
Played (reconstr, _) =>
|
blanchet@51200
|
645 |
if isar_proofs = SOME true then (maybe_isar_name, [])
|
blanchet@51200
|
646 |
else extract_reconstructor params reconstr
|
blanchet@51200
|
647 |
| _ => (maybe_isar_name, [])
|
blanchet@51200
|
648 |
in minimize_command override_params min_name end
|
blanchet@43051
|
649 |
|
blanchet@53480
|
650 |
val max_fact_instances = 10 (* FUDGE *)
|
blanchet@53480
|
651 |
|
blanchet@52034
|
652 |
fun repair_monomorph_context max_iters best_max_iters max_new_instances
|
blanchet@52034
|
653 |
best_max_new_instances =
|
blanchet@52034
|
654 |
Config.put Monomorph.max_rounds (max_iters |> the_default best_max_iters)
|
blanchet@52034
|
655 |
#> Config.put Monomorph.max_new_instances
|
blanchet@52034
|
656 |
(max_new_instances |> the_default best_max_new_instances)
|
blanchet@53480
|
657 |
#> Config.put Monomorph.max_thm_instances max_fact_instances
|
blanchet@52034
|
658 |
|
blanchet@51998
|
659 |
fun suffix_of_mode Auto_Try = "_try"
|
blanchet@51998
|
660 |
| suffix_of_mode Try = "_try"
|
blanchet@51998
|
661 |
| suffix_of_mode Normal = ""
|
blanchet@51998
|
662 |
| suffix_of_mode MaSh = ""
|
blanchet@51998
|
663 |
| suffix_of_mode Auto_Minimize = "_min"
|
blanchet@51998
|
664 |
| suffix_of_mode Minimize = "_min"
|
blanchet@44509
|
665 |
|
blanchet@44423
|
666 |
(* Give the ATPs some slack before interrupting them the hard way. "z3_tptp" on
|
blanchet@43631
|
667 |
Linux appears to be the only ATP that does not honor its time limit. *)
|
blanchet@43690
|
668 |
val atp_timeout_slack = seconds 1.0
|
blanchet@43631
|
669 |
|
blanchet@48331
|
670 |
val mono_max_privileged_facts = 10
|
blanchet@48331
|
671 |
|
blanchet@51186
|
672 |
(* For low values of "max_facts", this fudge value ensures that most slices are
|
blanchet@51186
|
673 |
invoked with a nontrivial amount of facts. *)
|
blanchet@51186
|
674 |
val max_fact_factor_fudge = 5
|
blanchet@51186
|
675 |
|
blanchet@43021
|
676 |
fun run_atp mode name
|
blanchet@48376
|
677 |
({exec, arguments, proof_delims, known_failures, prem_role, best_slices,
|
blanchet@48376
|
678 |
best_max_mono_iters, best_max_new_mono_instances, ...} : atp_config)
|
blanchet@46301
|
679 |
(params as {debug, verbose, overlord, type_enc, strict, lam_trans,
|
blanchet@51024
|
680 |
uncurried_aliases, fact_filter, max_facts, max_mono_iters,
|
smolkas@51130
|
681 |
max_new_mono_instances, isar_proofs, isar_compress,
|
blanchet@53764
|
682 |
isar_try0, slice, timeout, preplay_timeout, ...})
|
blanchet@43037
|
683 |
minimize_command
|
blanchet@51013
|
684 |
({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
|
blanchet@38023
|
685 |
let
|
blanchet@42182
|
686 |
val thy = Proof.theory_of state
|
blanchet@39318
|
687 |
val ctxt = Proof.context_of state
|
blanchet@47946
|
688 |
val atp_mode =
|
blanchet@48143
|
689 |
if Config.get ctxt completish then Sledgehammer_Completish
|
blanchet@47946
|
690 |
else Sledgehammer
|
blanchet@52196
|
691 |
val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt
|
blanchet@41159
|
692 |
val (dest_dir, problem_prefix) =
|
blanchet@51998
|
693 |
if overlord then overlord_file_location_of_prover name
|
blanchet@41159
|
694 |
else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix)
|
blanchet@40061
|
695 |
val problem_file_name =
|
blanchet@41159
|
696 |
Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^
|
blanchet@51998
|
697 |
suffix_of_mode mode ^ "_" ^ string_of_int subgoal)
|
blanchet@48656
|
698 |
val prob_path =
|
blanchet@40059
|
699 |
if dest_dir = "" then
|
blanchet@40061
|
700 |
File.tmp_path problem_file_name
|
blanchet@40059
|
701 |
else if File.exists (Path.explode dest_dir) then
|
blanchet@40061
|
702 |
Path.append (Path.explode dest_dir) problem_file_name
|
blanchet@39003
|
703 |
else
|
blanchet@40059
|
704 |
error ("No such directory: " ^ quote dest_dir ^ ".")
|
blanchet@52754
|
705 |
val exec = exec ()
|
blanchet@48376
|
706 |
val command0 =
|
blanchet@47055
|
707 |
case find_first (fn var => getenv var <> "") (fst exec) of
|
blanchet@48376
|
708 |
SOME var =>
|
blanchet@48376
|
709 |
let
|
blanchet@48376
|
710 |
val pref = getenv var ^ "/"
|
blanchet@48376
|
711 |
val paths = map (Path.explode o prefix pref) (snd exec)
|
blanchet@48376
|
712 |
in
|
blanchet@48376
|
713 |
case find_first File.exists paths of
|
blanchet@48376
|
714 |
SOME path => path
|
blanchet@48376
|
715 |
| NONE => error ("Bad executable: " ^ Path.print (hd paths) ^ ".")
|
blanchet@48376
|
716 |
end
|
blanchet@53989
|
717 |
| NONE => error ("The environment variable " ^ quote (List.last (fst exec)) ^
|
blanchet@47055
|
718 |
" is not set.")
|
blanchet@38023
|
719 |
fun split_time s =
|
blanchet@38023
|
720 |
let
|
blanchet@42448
|
721 |
val split = String.tokens (fn c => str c = "\n")
|
blanchet@47737
|
722 |
val (output, t) =
|
blanchet@47737
|
723 |
s |> split |> (try split_last #> the_default ([], "0"))
|
blanchet@47737
|
724 |
|>> cat_lines
|
blanchet@42448
|
725 |
fun as_num f = f >> (fst o read_int)
|
blanchet@42448
|
726 |
val num = as_num (Scan.many1 Symbol.is_ascii_digit)
|
blanchet@42448
|
727 |
val digit = Scan.one Symbol.is_ascii_digit
|
blanchet@42448
|
728 |
val num3 = as_num (digit ::: digit ::: (digit >> single))
|
blanchet@42448
|
729 |
val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b)
|
blanchet@45381
|
730 |
val as_time =
|
blanchet@45381
|
731 |
raw_explode #> Scan.read Symbol.stopper time #> the_default 0
|
blanchet@47737
|
732 |
in (output, as_time t |> Time.fromMilliseconds) end
|
blanchet@48656
|
733 |
fun run () =
|
blanchet@48376
|
734 |
let
|
blanchet@48376
|
735 |
(* If slicing is disabled, we expand the last slice to fill the entire
|
blanchet@48376
|
736 |
time available. *)
|
blanchet@48376
|
737 |
val actual_slices = get_slices slice (best_slices ctxt)
|
blanchet@48376
|
738 |
val num_actual_slices = length actual_slices
|
blanchet@51186
|
739 |
val max_fact_factor =
|
blanchet@51186
|
740 |
case max_facts of
|
blanchet@51186
|
741 |
NONE => 1.0
|
blanchet@51186
|
742 |
| SOME max =>
|
blanchet@51186
|
743 |
Real.fromInt max
|
blanchet@51186
|
744 |
/ Real.fromInt (fold (Integer.max o slice_max_facts)
|
blanchet@51186
|
745 |
actual_slices 0)
|
blanchet@48376
|
746 |
fun monomorphize_facts facts =
|
blanchet@48376
|
747 |
let
|
blanchet@48376
|
748 |
val ctxt =
|
blanchet@48376
|
749 |
ctxt
|
blanchet@53478
|
750 |
|> repair_monomorph_context max_mono_iters
|
blanchet@52034
|
751 |
best_max_mono_iters max_new_mono_instances
|
blanchet@52034
|
752 |
best_max_new_mono_instances
|
blanchet@48376
|
753 |
(* pseudo-theorem involving the same constants as the subgoal *)
|
blanchet@48376
|
754 |
val subgoal_th =
|
blanchet@48376
|
755 |
Logic.list_implies (hyp_ts, concl_t) |> Skip_Proof.make_thm thy
|
blanchet@48376
|
756 |
val rths =
|
blanchet@48376
|
757 |
facts |> chop mono_max_privileged_facts
|
blanchet@48376
|
758 |
|>> map (pair 1 o snd)
|
blanchet@48376
|
759 |
||> map (pair 2 o snd)
|
blanchet@48376
|
760 |
|> op @
|
blanchet@48376
|
761 |
|> cons (0, subgoal_th)
|
blanchet@48376
|
762 |
in
|
blanchet@53478
|
763 |
Monomorph.monomorph atp_schematic_consts_of ctxt rths
|
blanchet@53478
|
764 |
|> tl |> curry ListPair.zip (map fst facts)
|
blanchet@48376
|
765 |
|> maps (fn (name, rths) =>
|
blanchet@48376
|
766 |
map (pair name o zero_var_indexes o snd) rths)
|
blanchet@48376
|
767 |
end
|
blanchet@48376
|
768 |
fun run_slice time_left (cache_key, cache_value)
|
blanchet@48716
|
769 |
(slice, (time_frac,
|
blanchet@51011
|
770 |
(key as ((best_max_facts, best_fact_filter), format,
|
blanchet@51011
|
771 |
best_type_enc, best_lam_trans,
|
blanchet@51011
|
772 |
best_uncurried_aliases),
|
blanchet@48716
|
773 |
extra))) =
|
blanchet@38032
|
774 |
let
|
blanchet@51024
|
775 |
val effective_fact_filter =
|
blanchet@51024
|
776 |
fact_filter |> the_default best_fact_filter
|
blanchet@51998
|
777 |
val facts = get_facts_of_filter effective_fact_filter factss
|
blanchet@48376
|
778 |
val num_facts =
|
blanchet@51186
|
779 |
Real.ceil (max_fact_factor * Real.fromInt best_max_facts) +
|
blanchet@51186
|
780 |
max_fact_factor_fudge
|
blanchet@51186
|
781 |
|> Integer.min (length facts)
|
blanchet@48376
|
782 |
val soundness = if strict then Strict else Non_Strict
|
blanchet@48376
|
783 |
val type_enc =
|
blanchet@48376
|
784 |
type_enc |> choose_type_enc soundness best_type_enc format
|
blanchet@48376
|
785 |
val sound = is_type_enc_sound type_enc
|
blanchet@48376
|
786 |
val real_ms = Real.fromInt o Time.toMilliseconds
|
blanchet@48376
|
787 |
val slice_timeout =
|
blanchet@50557
|
788 |
case time_left of
|
blanchet@50557
|
789 |
SOME time_left =>
|
blanchet@50557
|
790 |
((real_ms time_left
|
blanchet@50557
|
791 |
|> (if slice < num_actual_slices - 1 then
|
blanchet@50557
|
792 |
curry Real.min (time_frac * real_ms (the timeout))
|
blanchet@50557
|
793 |
else
|
blanchet@50557
|
794 |
I))
|
blanchet@50557
|
795 |
* 0.001)
|
blanchet@50557
|
796 |
|> seconds |> SOME
|
blanchet@50557
|
797 |
| NONE => NONE
|
blanchet@48376
|
798 |
val generous_slice_timeout =
|
blanchet@50558
|
799 |
if mode = MaSh then NONE
|
blanchet@50558
|
800 |
else Option.map (curry Time.+ atp_timeout_slack) slice_timeout
|
blanchet@48376
|
801 |
val _ =
|
blanchet@48376
|
802 |
if debug then
|
blanchet@48376
|
803 |
quote name ^ " slice #" ^ string_of_int (slice + 1) ^
|
blanchet@48376
|
804 |
" with " ^ string_of_int num_facts ^ " fact" ^
|
blanchet@50557
|
805 |
plural_s num_facts ^
|
blanchet@50557
|
806 |
(case slice_timeout of
|
blanchet@52031
|
807 |
SOME timeout => " for " ^ string_of_time timeout
|
blanchet@50557
|
808 |
| NONE => "") ^ "..."
|
blanchet@48376
|
809 |
|> Output.urgent_message
|
blanchet@48376
|
810 |
else
|
blanchet@48376
|
811 |
()
|
blanchet@48376
|
812 |
val readable_names = not (Config.get ctxt atp_full_names)
|
blanchet@48376
|
813 |
val lam_trans =
|
blanchet@48376
|
814 |
case lam_trans of
|
blanchet@48376
|
815 |
SOME s => s
|
blanchet@48376
|
816 |
| NONE => best_lam_trans
|
blanchet@48376
|
817 |
val uncurried_aliases =
|
blanchet@48376
|
818 |
case uncurried_aliases of
|
blanchet@48376
|
819 |
SOME b => b
|
blanchet@48376
|
820 |
| NONE => best_uncurried_aliases
|
blanchet@48376
|
821 |
val value as (atp_problem, _, fact_names, _, _) =
|
blanchet@48376
|
822 |
if cache_key = SOME key then
|
blanchet@48376
|
823 |
cache_value
|
blanchet@48376
|
824 |
else
|
blanchet@48376
|
825 |
facts
|
blanchet@48376
|
826 |
|> not sound
|
blanchet@48376
|
827 |
? filter_out (is_dangerous_prop ctxt o prop_of o snd)
|
blanchet@48376
|
828 |
|> take num_facts
|
blanchet@48376
|
829 |
|> not (is_type_enc_polymorphic type_enc) ? monomorphize_facts
|
blanchet@48376
|
830 |
|> map (apsnd prop_of)
|
blanchet@48376
|
831 |
|> prepare_atp_problem ctxt format prem_role type_enc atp_mode
|
blanchet@48376
|
832 |
lam_trans uncurried_aliases
|
blanchet@48376
|
833 |
readable_names true hyp_ts concl_t
|
blanchet@48376
|
834 |
fun sel_weights () = atp_problem_selection_weights atp_problem
|
blanchet@48376
|
835 |
fun ord_info () = atp_problem_term_order_info atp_problem
|
blanchet@48376
|
836 |
val ord = effective_term_order ctxt name
|
blanchet@53492
|
837 |
val full_proof = isar_proofs |> the_default (mode = Minimize)
|
blanchet@50927
|
838 |
val args =
|
blanchet@50927
|
839 |
arguments ctxt full_proof extra
|
blanchet@50927
|
840 |
(slice_timeout |> the_default one_day)
|
blanchet@50927
|
841 |
(File.shell_path prob_path) (ord, ord_info, sel_weights)
|
blanchet@48376
|
842 |
val command =
|
blanchet@50927
|
843 |
"(exec 2>&1; " ^ File.shell_path command0 ^ " " ^ args ^ " " ^ ")"
|
blanchet@48532
|
844 |
|> enclose "TIMEFORMAT='%3R'; { time " " ; }"
|
blanchet@48376
|
845 |
val _ =
|
blanchet@48376
|
846 |
atp_problem
|
blanchet@51998
|
847 |
|> lines_of_atp_problem format ord ord_info
|
blanchet@48376
|
848 |
|> cons ("% " ^ command ^ "\n")
|
blanchet@48656
|
849 |
|> File.write_list prob_path
|
blanchet@52034
|
850 |
val ((output, run_time), (atp_proof, outcome)) =
|
blanchet@50557
|
851 |
time_limit generous_slice_timeout Isabelle_System.bash_output
|
blanchet@50557
|
852 |
command
|
blanchet@48376
|
853 |
|>> (if overlord then
|
blanchet@48376
|
854 |
prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
|
blanchet@48376
|
855 |
else
|
blanchet@48376
|
856 |
I)
|
blanchet@48376
|
857 |
|> fst |> split_time
|
blanchet@48376
|
858 |
|> (fn accum as (output, _) =>
|
blanchet@52034
|
859 |
(accum,
|
blanchet@48716
|
860 |
extract_tstplike_proof_and_outcome verbose proof_delims
|
blanchet@48716
|
861 |
known_failures output
|
blanchet@52031
|
862 |
|>> atp_proof_of_tstplike_proof atp_problem
|
blanchet@48376
|
863 |
handle UNRECOGNIZED_ATP_PROOF () =>
|
blanchet@48376
|
864 |
([], SOME ProofIncomplete)))
|
blanchet@48376
|
865 |
handle TimeLimit.TimeOut =>
|
blanchet@52034
|
866 |
(("", the slice_timeout), ([], SOME TimedOut))
|
blanchet@48376
|
867 |
val outcome =
|
blanchet@48376
|
868 |
case outcome of
|
blanchet@48376
|
869 |
NONE =>
|
blanchet@48376
|
870 |
(case used_facts_in_unsound_atp_proof ctxt fact_names atp_proof
|
blanchet@48376
|
871 |
|> Option.map (sort string_ord) of
|
blanchet@48376
|
872 |
SOME facts =>
|
blanchet@48376
|
873 |
let
|
blanchet@48376
|
874 |
val failure =
|
blanchet@48376
|
875 |
UnsoundProof (is_type_enc_sound type_enc, facts)
|
blanchet@48376
|
876 |
in
|
blanchet@53586
|
877 |
if debug then
|
blanchet@53586
|
878 |
(warning (string_of_atp_failure failure); NONE)
|
blanchet@53586
|
879 |
else
|
blanchet@53586
|
880 |
SOME failure
|
blanchet@48376
|
881 |
end
|
blanchet@48376
|
882 |
| NONE => NONE)
|
blanchet@48376
|
883 |
| _ => outcome
|
blanchet@51013
|
884 |
in
|
blanchet@51013
|
885 |
((SOME key, value), (output, run_time, facts, atp_proof, outcome))
|
blanchet@51013
|
886 |
end
|
blanchet@48376
|
887 |
val timer = Timer.startRealTimer ()
|
blanchet@48376
|
888 |
fun maybe_run_slice slice
|
blanchet@51013
|
889 |
(result as (cache, (_, run_time0, _, _, SOME _))) =
|
blanchet@48376
|
890 |
let
|
blanchet@50557
|
891 |
val time_left =
|
blanchet@50557
|
892 |
Option.map
|
blanchet@50557
|
893 |
(fn timeout => Time.- (timeout, Timer.checkRealTimer timer))
|
blanchet@50557
|
894 |
timeout
|
blanchet@48376
|
895 |
in
|
blanchet@50557
|
896 |
if time_left <> NONE andalso
|
blanchet@50557
|
897 |
Time.<= (the time_left, Time.zeroTime) then
|
blanchet@48376
|
898 |
result
|
blanchet@48376
|
899 |
else
|
blanchet@48376
|
900 |
run_slice time_left cache slice
|
blanchet@51013
|
901 |
|> (fn (cache, (output, run_time, used_from, atp_proof,
|
blanchet@51013
|
902 |
outcome)) =>
|
blanchet@51013
|
903 |
(cache, (output, Time.+ (run_time0, run_time), used_from,
|
blanchet@48376
|
904 |
atp_proof, outcome)))
|
blanchet@48376
|
905 |
end
|
blanchet@48376
|
906 |
| maybe_run_slice _ result = result
|
blanchet@48376
|
907 |
in
|
blanchet@48376
|
908 |
((NONE, ([], Symtab.empty, Vector.fromList [], [], Symtab.empty)),
|
blanchet@51013
|
909 |
("", Time.zeroTime, [], [], SOME InternalError))
|
blanchet@48376
|
910 |
|> fold maybe_run_slice actual_slices
|
blanchet@48376
|
911 |
end
|
blanchet@38023
|
912 |
(* If the problem file has not been exported, remove it; otherwise, export
|
blanchet@38023
|
913 |
the proof file too. *)
|
blanchet@48656
|
914 |
fun clean_up () =
|
blanchet@48656
|
915 |
if dest_dir = "" then (try File.rm prob_path; ()) else ()
|
blanchet@51013
|
916 |
fun export (_, (output, _, _, _, _)) =
|
blanchet@48376
|
917 |
if dest_dir = "" then ()
|
blanchet@48656
|
918 |
else File.write (Path.explode (Path.implode prob_path ^ "_proof")) output
|
blanchet@52150
|
919 |
val ((_, (_, pool, fact_names, lifted, sym_tab)),
|
blanchet@51013
|
920 |
(output, run_time, used_from, atp_proof, outcome)) =
|
blanchet@48656
|
921 |
with_cleanup clean_up run () |> tap export
|
blanchet@39492
|
922 |
val important_message =
|
wenzelm@53052
|
923 |
if mode = Normal andalso random_range 0 (atp_important_message_keep_quotient - 1) = 0
|
wenzelm@53052
|
924 |
then
|
blanchet@39492
|
925 |
extract_important_message output
|
blanchet@39492
|
926 |
else
|
blanchet@39492
|
927 |
""
|
blanchet@43261
|
928 |
val (used_facts, preplay, message, message_tail) =
|
blanchet@38023
|
929 |
case outcome of
|
blanchet@38023
|
930 |
NONE =>
|
blanchet@43033
|
931 |
let
|
blanchet@45551
|
932 |
val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof
|
blanchet@45590
|
933 |
val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof
|
blanchet@45521
|
934 |
val reconstrs =
|
blanchet@45554
|
935 |
bunch_of_reconstructors needs_full_types
|
blanchet@52031
|
936 |
(lam_trans_of_atp_proof atp_proof
|
blanchet@46405
|
937 |
o (fn desperate => if desperate then hide_lamsN
|
blanchet@46405
|
938 |
else metis_default_lam_trans))
|
blanchet@43033
|
939 |
in
|
blanchet@43052
|
940 |
(used_facts,
|
blanchet@50669
|
941 |
Lazy.lazy (fn () =>
|
blanchet@50669
|
942 |
let
|
blanchet@50669
|
943 |
val used_pairs =
|
blanchet@51013
|
944 |
used_from |> filter_used_facts false used_facts
|
blanchet@50669
|
945 |
in
|
blanchet@50669
|
946 |
play_one_line_proof mode debug verbose preplay_timeout
|
blanchet@50669
|
947 |
used_pairs state subgoal (hd reconstrs) reconstrs
|
blanchet@50669
|
948 |
end),
|
blanchet@43052
|
949 |
fn preplay =>
|
blanchet@43052
|
950 |
let
|
blanchet@49921
|
951 |
val _ =
|
blanchet@49921
|
952 |
if verbose then
|
blanchet@51232
|
953 |
Output.urgent_message "Generating proof text..."
|
blanchet@49921
|
954 |
else
|
blanchet@49921
|
955 |
()
|
blanchet@43052
|
956 |
val isar_params =
|
blanchet@53763
|
957 |
(debug, verbose, preplay_timeout, isar_compress, isar_try0,
|
blanchet@53764
|
958 |
pool, fact_names, lifted, sym_tab, atp_proof, goal)
|
blanchet@43052
|
959 |
val one_line_params =
|
blanchet@43052
|
960 |
(preplay, proof_banner mode name, used_facts,
|
blanchet@51200
|
961 |
choose_minimize_command ctxt params minimize_command name
|
blanchet@51200
|
962 |
preplay,
|
blanchet@43052
|
963 |
subgoal, subgoal_count)
|
blanchet@48799
|
964 |
val num_chained = length (#facts (Proof.goal state))
|
blanchet@48799
|
965 |
in
|
wenzelm@53052
|
966 |
proof_text ctxt isar_proofs isar_params
|
blanchet@52754
|
967 |
num_chained one_line_params
|
blanchet@48799
|
968 |
end,
|
blanchet@43261
|
969 |
(if verbose then
|
blanchet@52031
|
970 |
"\nATP real CPU time: " ^ string_of_time run_time ^ "."
|
blanchet@43261
|
971 |
else
|
blanchet@43261
|
972 |
"") ^
|
blanchet@43261
|
973 |
(if important_message <> "" then
|
blanchet@43261
|
974 |
"\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^
|
blanchet@43261
|
975 |
important_message
|
blanchet@43261
|
976 |
else
|
blanchet@43261
|
977 |
""))
|
blanchet@43033
|
978 |
end
|
blanchet@43052
|
979 |
| SOME failure =>
|
blanchet@50669
|
980 |
([], Lazy.value (Failed_to_Play plain_metis),
|
blanchet@53586
|
981 |
fn _ => string_of_atp_failure failure, "")
|
blanchet@38023
|
982 |
in
|
blanchet@51013
|
983 |
{outcome = outcome, used_facts = used_facts, used_from = used_from,
|
blanchet@51009
|
984 |
run_time = run_time, preplay = preplay, message = message,
|
blanchet@51009
|
985 |
message_tail = message_tail}
|
blanchet@38023
|
986 |
end
|
blanchet@38023
|
987 |
|
blanchet@51014
|
988 |
fun rotate_one (x :: xs) = xs @ [x]
|
blanchet@51014
|
989 |
|
blanchet@40669
|
990 |
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until
|
blanchet@40669
|
991 |
these are sorted out properly in the SMT module, we have to interpret these
|
blanchet@40669
|
992 |
ourselves. *)
|
blanchet@40684
|
993 |
val remote_smt_failures =
|
blanchet@43631
|
994 |
[(2, NoLibwwwPerl),
|
blanchet@43631
|
995 |
(22, CantConnect)]
|
blanchet@40684
|
996 |
val z3_failures =
|
blanchet@41236
|
997 |
[(101, OutOfResources),
|
blanchet@41236
|
998 |
(103, MalformedInput),
|
blanchet@50667
|
999 |
(110, MalformedInput),
|
blanchet@50667
|
1000 |
(112, TimedOut)]
|
blanchet@40684
|
1001 |
val unix_failures =
|
blanchet@48797
|
1002 |
[(138, Crashed),
|
blanchet@48797
|
1003 |
(139, Crashed)]
|
blanchet@43631
|
1004 |
val smt_failures = remote_smt_failures @ z3_failures @ unix_failures
|
blanchet@40555
|
1005 |
|
blanchet@52031
|
1006 |
fun failure_of_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
|
blanchet@43050
|
1007 |
if is_real_cex then Unprovable else GaveUp
|
blanchet@52031
|
1008 |
| failure_of_smt_failure SMT_Failure.Time_Out = TimedOut
|
blanchet@52031
|
1009 |
| failure_of_smt_failure (SMT_Failure.Abnormal_Termination code) =
|
blanchet@41222
|
1010 |
(case AList.lookup (op =) smt_failures code of
|
blanchet@40684
|
1011 |
SOME failure => failure
|
blanchet@41259
|
1012 |
| NONE => UnknownError ("Abnormal termination with exit code " ^
|
blanchet@41259
|
1013 |
string_of_int code ^ "."))
|
blanchet@52031
|
1014 |
| failure_of_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources
|
blanchet@52031
|
1015 |
| failure_of_smt_failure (SMT_Failure.Other_Failure s) = UnknownError s
|
blanchet@40063
|
1016 |
|
blanchet@40698
|
1017 |
(* FUDGE *)
|
blanchet@42646
|
1018 |
val smt_max_slices =
|
blanchet@42646
|
1019 |
Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
|
blanchet@42646
|
1020 |
val smt_slice_fact_frac =
|
blanchet@51014
|
1021 |
Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac}
|
blanchet@51014
|
1022 |
(K 0.667)
|
blanchet@42646
|
1023 |
val smt_slice_time_frac =
|
blanchet@51014
|
1024 |
Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.333)
|
blanchet@42646
|
1025 |
val smt_slice_min_secs =
|
blanchet@51014
|
1026 |
Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 3)
|
blanchet@40409
|
1027 |
|
blanchet@54000
|
1028 |
val is_boring_builtin_typ =
|
blanchet@54000
|
1029 |
not o exists_subtype (member (op =) [@{typ nat}, @{typ int}, HOLogic.realT])
|
blanchet@54000
|
1030 |
|
blanchet@50759
|
1031 |
fun smt_filter_loop name
|
blanchet@42724
|
1032 |
({debug, verbose, overlord, max_mono_iters,
|
blanchet@45706
|
1033 |
max_new_mono_instances, timeout, slice, ...} : params)
|
blanchet@50486
|
1034 |
state goal i =
|
blanchet@40409
|
1035 |
let
|
blanchet@50759
|
1036 |
fun repair_context ctxt =
|
blanchet@50759
|
1037 |
ctxt |> select_smt_solver name
|
blanchet@50759
|
1038 |
|> Config.put SMT_Config.verbose debug
|
blanchet@50759
|
1039 |
|> (if overlord then
|
blanchet@50759
|
1040 |
Config.put SMT_Config.debug_files
|
blanchet@51998
|
1041 |
(overlord_file_location_of_prover name
|
blanchet@50759
|
1042 |
|> (fn (path, name) => path ^ "/" ^ name))
|
blanchet@50759
|
1043 |
else
|
blanchet@50759
|
1044 |
I)
|
blanchet@50759
|
1045 |
|> Config.put SMT_Config.infer_triggers
|
blanchet@50759
|
1046 |
(Config.get ctxt smt_triggers)
|
blanchet@54000
|
1047 |
|> not (Config.get ctxt smt_builtin_trans)
|
blanchet@54000
|
1048 |
? (SMT_Builtin.filter_builtins is_boring_builtin_typ
|
blanchet@54000
|
1049 |
#> Config.put SMT_Config.datatypes false)
|
blanchet@52034
|
1050 |
|> repair_monomorph_context max_mono_iters default_max_mono_iters
|
blanchet@52034
|
1051 |
max_new_mono_instances default_max_new_mono_instances
|
blanchet@52034
|
1052 |
val state = Proof.map_context (repair_context) state
|
blanchet@52034
|
1053 |
val ctxt = Proof.context_of state
|
blanchet@45706
|
1054 |
val max_slices = if slice then Config.get ctxt smt_max_slices else 1
|
blanchet@51024
|
1055 |
fun do_slice timeout slice outcome0 time_so_far
|
blanchet@51024
|
1056 |
(weighted_factss as (fact_filter, weighted_facts) :: _) =
|
blanchet@40553
|
1057 |
let
|
blanchet@40553
|
1058 |
val timer = Timer.startRealTimer ()
|
blanchet@42443
|
1059 |
val slice_timeout =
|
blanchet@50557
|
1060 |
if slice < max_slices andalso timeout <> NONE then
|
blanchet@50557
|
1061 |
let val ms = timeout |> the |> Time.toMilliseconds in
|
blanchet@50557
|
1062 |
Int.min (ms,
|
blanchet@50557
|
1063 |
Int.max (1000 * Config.get ctxt smt_slice_min_secs,
|
blanchet@50557
|
1064 |
Real.ceil (Config.get ctxt smt_slice_time_frac
|
blanchet@50557
|
1065 |
* Real.fromInt ms)))
|
blanchet@50557
|
1066 |
|> Time.fromMilliseconds |> SOME
|
blanchet@50557
|
1067 |
end
|
blanchet@40553
|
1068 |
else
|
blanchet@40553
|
1069 |
timeout
|
blanchet@51009
|
1070 |
val num_facts = length weighted_facts
|
blanchet@40553
|
1071 |
val _ =
|
blanchet@42614
|
1072 |
if debug then
|
blanchet@42614
|
1073 |
quote name ^ " slice " ^ string_of_int slice ^ " with " ^
|
blanchet@50557
|
1074 |
string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
|
blanchet@50557
|
1075 |
(case slice_timeout of
|
blanchet@52031
|
1076 |
SOME timeout => " for " ^ string_of_time timeout
|
blanchet@50557
|
1077 |
| NONE => "") ^ "..."
|
blanchet@40553
|
1078 |
|> Output.urgent_message
|
blanchet@40553
|
1079 |
else
|
blanchet@40553
|
1080 |
()
|
blanchet@41168
|
1081 |
val birth = Timer.checkRealTimer timer
|
blanchet@41171
|
1082 |
val _ =
|
blanchet@41211
|
1083 |
if debug then Output.urgent_message "Invoking SMT solver..." else ()
|
blanchet@41209
|
1084 |
val (outcome, used_facts) =
|
blanchet@51204
|
1085 |
SMT_Solver.smt_filter_preprocess ctxt [] goal weighted_facts i
|
blanchet@50557
|
1086 |
|> SMT_Solver.smt_filter_apply (slice_timeout |> the_default one_day)
|
blanchet@41239
|
1087 |
|> (fn {outcome, used_facts} => (outcome, used_facts))
|
blanchet@41209
|
1088 |
handle exn => if Exn.is_interrupt exn then
|
blanchet@41209
|
1089 |
reraise exn
|
blanchet@41209
|
1090 |
else
|
blanchet@42061
|
1091 |
(ML_Compiler.exn_message exn
|
blanchet@41209
|
1092 |
|> SMT_Failure.Other_Failure |> SOME, [])
|
blanchet@41168
|
1093 |
val death = Timer.checkRealTimer timer
|
blanchet@40553
|
1094 |
val outcome0 = if is_none outcome0 then SOME outcome else outcome0
|
blanchet@41168
|
1095 |
val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
|
blanchet@40553
|
1096 |
val too_many_facts_perhaps =
|
blanchet@40553
|
1097 |
case outcome of
|
blanchet@40553
|
1098 |
NONE => false
|
blanchet@40553
|
1099 |
| SOME (SMT_Failure.Counterexample _) => false
|
blanchet@42443
|
1100 |
| SOME SMT_Failure.Time_Out => slice_timeout <> timeout
|
blanchet@42614
|
1101 |
| SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *)
|
blanchet@40553
|
1102 |
| SOME SMT_Failure.Out_Of_Memory => true
|
blanchet@41211
|
1103 |
| SOME (SMT_Failure.Other_Failure _) => true
|
blanchet@50557
|
1104 |
val timeout =
|
blanchet@50557
|
1105 |
Option.map
|
blanchet@50557
|
1106 |
(fn timeout => Time.- (timeout, Timer.checkRealTimer timer))
|
blanchet@50557
|
1107 |
timeout
|
blanchet@40553
|
1108 |
in
|
blanchet@42443
|
1109 |
if too_many_facts_perhaps andalso slice < max_slices andalso
|
blanchet@50557
|
1110 |
num_facts > 0 andalso
|
blanchet@50557
|
1111 |
(timeout = NONE orelse Time.> (the timeout, Time.zeroTime)) then
|
blanchet@41169
|
1112 |
let
|
blanchet@42614
|
1113 |
val new_num_facts =
|
blanchet@42646
|
1114 |
Real.ceil (Config.get ctxt smt_slice_fact_frac
|
blanchet@42646
|
1115 |
* Real.fromInt num_facts)
|
blanchet@51024
|
1116 |
val weighted_factss as (new_fact_filter, _) :: _ =
|
blanchet@51024
|
1117 |
weighted_factss
|
blanchet@51024
|
1118 |
|> rotate_one
|
blanchet@51024
|
1119 |
|> app_hd (apsnd (take new_num_facts))
|
blanchet@51024
|
1120 |
val show_filter = fact_filter <> new_fact_filter
|
blanchet@51024
|
1121 |
fun num_of_facts fact_filter num_facts =
|
blanchet@51024
|
1122 |
string_of_int num_facts ^
|
blanchet@51024
|
1123 |
(if show_filter then " " ^ quote fact_filter else "") ^
|
blanchet@51024
|
1124 |
" fact" ^ plural_s num_facts
|
blanchet@42614
|
1125 |
val _ =
|
blanchet@53500
|
1126 |
if debug then
|
blanchet@51024
|
1127 |
quote name ^ " invoked with " ^
|
blanchet@51024
|
1128 |
num_of_facts fact_filter num_facts ^ ": " ^
|
blanchet@53586
|
1129 |
string_of_atp_failure (failure_of_smt_failure (the outcome)) ^
|
blanchet@51024
|
1130 |
" Retrying with " ^ num_of_facts new_fact_filter new_num_facts ^
|
blanchet@51024
|
1131 |
"..."
|
blanchet@42614
|
1132 |
|> Output.urgent_message
|
blanchet@42614
|
1133 |
else
|
blanchet@42614
|
1134 |
()
|
blanchet@42443
|
1135 |
in
|
blanchet@51024
|
1136 |
do_slice timeout (slice + 1) outcome0 time_so_far weighted_factss
|
blanchet@42443
|
1137 |
end
|
blanchet@40553
|
1138 |
else
|
blanchet@40553
|
1139 |
{outcome = if is_none outcome then NONE else the outcome0,
|
blanchet@51009
|
1140 |
used_facts = used_facts, used_from = map (apsnd snd) weighted_facts,
|
blanchet@51009
|
1141 |
run_time = time_so_far}
|
blanchet@40409
|
1142 |
end
|
blanchet@42443
|
1143 |
in do_slice timeout 1 NONE Time.zeroTime end
|
blanchet@40409
|
1144 |
|
blanchet@43052
|
1145 |
fun run_smt_solver mode name (params as {debug, verbose, preplay_timeout, ...})
|
blanchet@43011
|
1146 |
minimize_command
|
blanchet@51014
|
1147 |
({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
|
blanchet@36379
|
1148 |
let
|
blanchet@41242
|
1149 |
val ctxt = Proof.context_of state
|
blanchet@51014
|
1150 |
fun weight_facts facts =
|
blanchet@51014
|
1151 |
let val num_facts = length facts in
|
blanchet@51014
|
1152 |
facts ~~ (0 upto num_facts - 1)
|
blanchet@51014
|
1153 |
|> map (weight_smt_fact ctxt num_facts)
|
blanchet@51014
|
1154 |
end
|
blanchet@51024
|
1155 |
val weighted_factss = factss |> map (apsnd weight_facts)
|
blanchet@51009
|
1156 |
val {outcome, used_facts = used_pairs, used_from, run_time} =
|
blanchet@51014
|
1157 |
smt_filter_loop name params state goal subgoal weighted_factss
|
blanchet@45781
|
1158 |
val used_facts = used_pairs |> map fst
|
blanchet@52031
|
1159 |
val outcome = outcome |> Option.map failure_of_smt_failure
|
blanchet@43261
|
1160 |
val (preplay, message, message_tail) =
|
blanchet@40184
|
1161 |
case outcome of
|
blanchet@40184
|
1162 |
NONE =>
|
blanchet@50669
|
1163 |
(Lazy.lazy (fn () =>
|
blanchet@50669
|
1164 |
play_one_line_proof mode debug verbose preplay_timeout used_pairs
|
blanchet@50669
|
1165 |
state subgoal SMT
|
blanchet@50669
|
1166 |
(bunch_of_reconstructors false (fn desperate =>
|
blanchet@50669
|
1167 |
if desperate then liftingN else metis_default_lam_trans))),
|
blanchet@43052
|
1168 |
fn preplay =>
|
blanchet@43052
|
1169 |
let
|
blanchet@43052
|
1170 |
val one_line_params =
|
blanchet@43052
|
1171 |
(preplay, proof_banner mode name, used_facts,
|
blanchet@51200
|
1172 |
choose_minimize_command ctxt params minimize_command name
|
blanchet@51200
|
1173 |
preplay,
|
blanchet@43052
|
1174 |
subgoal, subgoal_count)
|
blanchet@48799
|
1175 |
val num_chained = length (#facts (Proof.goal state))
|
blanchet@52754
|
1176 |
in
|
wenzelm@53052
|
1177 |
one_line_proof_text num_chained one_line_params
|
blanchet@52754
|
1178 |
end,
|
blanchet@43261
|
1179 |
if verbose then
|
blanchet@52031
|
1180 |
"\nSMT solver real CPU time: " ^ string_of_time run_time ^ "."
|
blanchet@43261
|
1181 |
else
|
blanchet@43261
|
1182 |
"")
|
blanchet@43166
|
1183 |
| SOME failure =>
|
blanchet@50669
|
1184 |
(Lazy.value (Failed_to_Play plain_metis),
|
blanchet@53586
|
1185 |
fn _ => string_of_atp_failure failure, "")
|
blanchet@40063
|
1186 |
in
|
blanchet@51009
|
1187 |
{outcome = outcome, used_facts = used_facts, used_from = used_from,
|
blanchet@51009
|
1188 |
run_time = run_time, preplay = preplay, message = message,
|
blanchet@51009
|
1189 |
message_tail = message_tail}
|
blanchet@40063
|
1190 |
end
|
blanchet@40063
|
1191 |
|
blanchet@45520
|
1192 |
fun run_reconstructor mode name
|
blanchet@45561
|
1193 |
(params as {debug, verbose, timeout, type_enc, lam_trans, ...})
|
blanchet@45379
|
1194 |
minimize_command
|
blanchet@51010
|
1195 |
({state, subgoal, subgoal_count, factss = (_, facts) :: _, ...}
|
blanchet@51007
|
1196 |
: prover_problem) =
|
blanchet@43050
|
1197 |
let
|
blanchet@45520
|
1198 |
val reconstr =
|
blanchet@45520
|
1199 |
if name = metisN then
|
blanchet@45520
|
1200 |
Metis (type_enc |> the_default (hd partial_type_encs),
|
blanchet@45520
|
1201 |
lam_trans |> the_default metis_default_lam_trans)
|
blanchet@45520
|
1202 |
else if name = smtN then
|
blanchet@45520
|
1203 |
SMT
|
blanchet@45520
|
1204 |
else
|
blanchet@45520
|
1205 |
raise Fail ("unknown reconstructor: " ^ quote name)
|
blanchet@51005
|
1206 |
val used_facts = facts |> map fst
|
blanchet@43050
|
1207 |
in
|
blanchet@45379
|
1208 |
case play_one_line_proof (if mode = Minimize then Normal else mode) debug
|
blanchet@51005
|
1209 |
verbose timeout facts state subgoal reconstr
|
blanchet@45520
|
1210 |
[reconstr] of
|
blanchet@43050
|
1211 |
play as Played (_, time) =>
|
blanchet@51009
|
1212 |
{outcome = NONE, used_facts = used_facts, used_from = facts,
|
blanchet@51009
|
1213 |
run_time = time, preplay = Lazy.value play,
|
blanchet@45561
|
1214 |
message =
|
blanchet@45561
|
1215 |
fn play =>
|
blanchet@45561
|
1216 |
let
|
blanchet@45561
|
1217 |
val (_, override_params) = extract_reconstructor params reconstr
|
blanchet@45561
|
1218 |
val one_line_params =
|
blanchet@45561
|
1219 |
(play, proof_banner mode name, used_facts,
|
blanchet@45561
|
1220 |
minimize_command override_params name, subgoal,
|
blanchet@45561
|
1221 |
subgoal_count)
|
blanchet@48799
|
1222 |
val num_chained = length (#facts (Proof.goal state))
|
blanchet@52754
|
1223 |
in
|
wenzelm@53052
|
1224 |
one_line_proof_text num_chained one_line_params
|
blanchet@52754
|
1225 |
end,
|
blanchet@43261
|
1226 |
message_tail = ""}
|
blanchet@43052
|
1227 |
| play =>
|
blanchet@43166
|
1228 |
let
|
blanchet@43166
|
1229 |
val failure = case play of Failed_to_Play _ => GaveUp | _ => TimedOut
|
blanchet@43166
|
1230 |
in
|
blanchet@51009
|
1231 |
{outcome = SOME failure, used_facts = [], used_from = [],
|
blanchet@51009
|
1232 |
run_time = Time.zeroTime, preplay = Lazy.value play,
|
blanchet@53586
|
1233 |
message = fn _ => string_of_atp_failure failure, message_tail = ""}
|
blanchet@43050
|
1234 |
end
|
blanchet@43050
|
1235 |
end
|
blanchet@43050
|
1236 |
|
blanchet@43021
|
1237 |
fun get_prover ctxt mode name =
|
wenzelm@42361
|
1238 |
let val thy = Proof_Context.theory_of ctxt in
|
blanchet@45379
|
1239 |
if is_reconstructor name then run_reconstructor mode name
|
blanchet@47606
|
1240 |
else if is_atp thy name then run_atp mode name (get_atp thy name ())
|
blanchet@43052
|
1241 |
else if is_smt_prover ctxt name then run_smt_solver mode name
|
blanchet@43052
|
1242 |
else error ("No such prover: " ^ name ^ ".")
|
blanchet@40941
|
1243 |
end
|
blanchet@40063
|
1244 |
|
wenzelm@28582
|
1245 |
end;
|