| author | blanchet | 
| Mon, 23 Jul 2012 15:32:30 +0200 | |
| changeset 48442 | 3c9890c19e90 | 
| parent 48392 | ca998fa08cd9 | 
| child 48532 | c0f44941e674 | 
| permissions | -rw-r--r-- | 
| 41087 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
41066diff
changeset | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_provers.ML | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 2 | Author: Fabian Immler, TU Muenchen | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 3 | Author: Makarius | 
| 35969 | 4 | Author: Jasmin Blanchette, TU Muenchen | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 5 | |
| 41087 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
41066diff
changeset | 6 | Generic prover abstraction for Sledgehammer. | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 7 | *) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 8 | |
| 41087 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
41066diff
changeset | 9 | signature SLEDGEHAMMER_PROVERS = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 10 | sig | 
| 40181 | 11 | type failure = ATP_Proof.failure | 
| 46340 | 12 | type stature = ATP_Problem_Generate.stature | 
| 46320 | 13 | type type_enc = ATP_Problem_Generate.type_enc | 
| 14 | type reconstructor = ATP_Proof_Reconstruct.reconstructor | |
| 15 | type play = ATP_Proof_Reconstruct.play | |
| 16 | type minimize_command = ATP_Proof_Reconstruct.minimize_command | |
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39492diff
changeset | 17 | |
| 48392 
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
 blanchet parents: 
48376diff
changeset | 18 | datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize | 
| 43021 | 19 | |
| 35969 | 20 | type params = | 
| 48321 | 21 |     {debug : bool,
 | 
| 22 | verbose : bool, | |
| 23 | overlord : bool, | |
| 24 | blocking : bool, | |
| 25 | provers : string list, | |
| 26 | type_enc : string option, | |
| 27 | strict : bool, | |
| 28 | lam_trans : string option, | |
| 29 | uncurried_aliases : bool option, | |
| 30 | learn : bool, | |
| 31 | fact_filter : string option, | |
| 32 | max_facts : int option, | |
| 33 | fact_thresholds : real * real, | |
| 34 | max_mono_iters : int option, | |
| 35 | max_new_mono_instances : int option, | |
| 36 | isar_proof : bool, | |
| 37 | isar_shrink_factor : int, | |
| 38 | slice : bool, | |
| 39 | minimize : bool option, | |
| 40 | timeout : Time.time, | |
| 41 | preplay_timeout : Time.time, | |
| 42 | expect : string} | |
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39492diff
changeset | 43 | |
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 44 | type relevance_fudge = | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 45 |     {local_const_multiplier : real,
 | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 46 | worse_irrel_freq : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 47 | higher_order_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 48 | abs_rel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 49 | abs_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 50 | skolem_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 51 | theory_const_rel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 52 | theory_const_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 53 | chained_const_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 54 | intro_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 55 | elim_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 56 | simp_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 57 | local_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 58 | assum_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 59 | chained_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 60 | max_imperfect : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 61 | max_imperfect_exp : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 62 | threshold_divisor : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 63 | ridiculous_threshold : real} | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 64 | |
| 41090 | 65 | datatype prover_fact = | 
| 46340 | 66 | Untranslated_Fact of (string * stature) * thm | | 
| 67 | SMT_Weighted_Fact of (string * stature) * (int option * thm) | |
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 68 | |
| 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 69 | type prover_problem = | 
| 48321 | 70 |     {state : Proof.state,
 | 
| 71 | goal : thm, | |
| 72 | subgoal : int, | |
| 73 | subgoal_count : int, | |
| 74 | facts : prover_fact list} | |
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39492diff
changeset | 75 | |
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 76 | type prover_result = | 
| 48321 | 77 |     {outcome : failure option,
 | 
| 78 | used_facts : (string * stature) list, | |
| 79 | run_time : Time.time, | |
| 80 | preplay : unit -> play, | |
| 81 | message : play -> string, | |
| 82 | message_tail : string} | |
| 39493 
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
 blanchet parents: 
39492diff
changeset | 83 | |
| 43051 | 84 | type prover = | 
| 45520 | 85 | params -> ((string * string list) list -> string -> minimize_command) | 
| 86 | -> prover_problem -> prover_result | |
| 35867 | 87 | |
| 43092 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 blanchet parents: 
43085diff
changeset | 88 | val dest_dir : string Config.T | 
| 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 blanchet parents: 
43085diff
changeset | 89 | val problem_prefix : string Config.T | 
| 48143 | 90 | val completish : bool Config.T | 
| 44592 
54906b0337ab
flip logic of boolean option so it's off by default
 blanchet parents: 
44586diff
changeset | 91 | val atp_full_names : bool Config.T | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 92 | val smt_triggers : bool Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 93 | val smt_weights : bool Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 94 | val smt_weight_min_facts : int Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 95 | val smt_min_weight : int Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 96 | val smt_max_weight : int Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 97 | val smt_max_weight_index : int Config.T | 
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 98 | val smt_weight_curve : (int -> int) Unsynchronized.ref | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 99 | val smt_max_slices : int Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 100 | val smt_slice_fact_frac : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 101 | val smt_slice_time_frac : real Config.T | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 102 | val smt_slice_min_secs : int Config.T | 
| 48319 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 blanchet parents: 
48314diff
changeset | 103 | val SledgehammerN : string | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 104 | val plain_metis : reconstructor | 
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 105 | val select_smt_solver : string -> Proof.context -> Proof.context | 
| 45520 | 106 | val extract_reconstructor : | 
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 107 | params -> reconstructor -> string * (string * string list) list | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 108 | val is_reconstructor : string -> bool | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 109 | val is_atp : theory -> string -> bool | 
| 41087 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
41066diff
changeset | 110 | val is_smt_prover : Proof.context -> string -> bool | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47946diff
changeset | 111 | val is_ho_atp: Proof.context -> string -> bool | 
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 112 | val is_unit_equational_atp : Proof.context -> string -> bool | 
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41723diff
changeset | 113 | val is_prover_supported : Proof.context -> string -> bool | 
| 40072 
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
 blanchet parents: 
40071diff
changeset | 114 | val is_prover_installed : Proof.context -> string -> bool | 
| 48293 | 115 | val default_max_facts_for_prover : Proof.context -> bool -> string -> int | 
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 116 | val is_unit_equality : term -> bool | 
| 42952 
96f62b77748f
tuning -- the "appropriate" terminology is inspired from TPTP
 blanchet parents: 
42944diff
changeset | 117 | val is_appropriate_prop_for_prover : Proof.context -> string -> term -> bool | 
| 40369 
53dca3bd4250
use the SMT integration's official list of built-ins
 blanchet parents: 
40341diff
changeset | 118 | val is_built_in_const_for_prover : | 
| 41336 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 119 | Proof.context -> string -> string * typ -> term list -> bool * term list | 
| 41087 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
41066diff
changeset | 120 | val atp_relevance_fudge : relevance_fudge | 
| 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
41066diff
changeset | 121 | val smt_relevance_fudge : relevance_fudge | 
| 40941 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 blanchet parents: 
40723diff
changeset | 122 | val relevance_fudge_for_prover : Proof.context -> string -> relevance_fudge | 
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 123 | val weight_smt_fact : | 
| 46340 | 124 | Proof.context -> int -> ((string * stature) * thm) * int | 
| 125 | -> (string * stature) * (int option * thm) | |
| 126 | val untranslated_fact : prover_fact -> (string * stature) * thm | |
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 127 | val smt_weighted_fact : | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 128 | Proof.context -> int -> prover_fact * int | 
| 46340 | 129 | -> (string * stature) * (int option * thm) | 
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41723diff
changeset | 130 | val supported_provers : Proof.context -> unit | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 131 | val kill_provers : unit -> unit | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 132 | val running_provers : unit -> unit | 
| 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 133 | val messages : int option -> unit | 
| 43033 | 134 |   val filter_used_facts : ''a list -> (''a * 'b) list -> (''a * 'b) list
 | 
| 43021 | 135 | val get_prover : Proof.context -> mode -> string -> prover | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 136 | end; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 137 | |
| 41087 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
41066diff
changeset | 138 | structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 139 | struct | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 140 | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43063diff
changeset | 141 | open ATP_Util | 
| 38028 | 142 | open ATP_Problem | 
| 39491 
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
 blanchet parents: 
39453diff
changeset | 143 | open ATP_Proof | 
| 38028 | 144 | open ATP_Systems | 
| 46320 | 145 | open ATP_Problem_Generate | 
| 146 | open ATP_Proof_Reconstruct | |
| 45521 | 147 | open Metis_Tactic | 
| 38023 | 148 | open Sledgehammer_Util | 
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 149 | |
| 37583 
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
 blanchet parents: 
37581diff
changeset | 150 | |
| 
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
 blanchet parents: 
37581diff
changeset | 151 | (** The Sledgehammer **) | 
| 
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
 blanchet parents: 
37581diff
changeset | 152 | |
| 48392 
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
 blanchet parents: 
48376diff
changeset | 153 | datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize | 
| 43021 | 154 | |
| 45376 | 155 | (* Identifier that distinguishes Sledgehammer from other tools that could use | 
| 38102 
019a49759829
fix bug in the newly introduced "bound concealing" code
 blanchet parents: 
38100diff
changeset | 156 | "Async_Manager". *) | 
| 48319 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 blanchet parents: 
48314diff
changeset | 157 | val SledgehammerN = "Sledgehammer" | 
| 37585 | 158 | |
| 45520 | 159 | val reconstructor_names = [metisN, smtN] | 
| 46365 | 160 | val plain_metis = Metis (hd partial_type_encs, combsN) | 
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 161 | val is_reconstructor = member (op =) reconstructor_names | 
| 43228 
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
 blanchet parents: 
43226diff
changeset | 162 | |
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 163 | val is_atp = member (op =) o supported_atps | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 164 | |
| 43233 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 165 | val select_smt_solver = Context.proof_map o SMT_Config.select_solver | 
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 166 | |
| 45376 | 167 | fun is_smt_prover ctxt = member (op =) (SMT_Solver.available_solvers_of ctxt) | 
| 40062 | 168 | |
| 44597 | 169 | fun is_atp_for_format is_format ctxt name = | 
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 170 | let val thy = Proof_Context.theory_of ctxt in | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 171 | case try (get_atp thy) name of | 
| 47606 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 blanchet parents: 
47531diff
changeset | 172 | SOME config => | 
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 173 | exists (fn (_, (_, ((_, format, _, _, _), _))) => is_format format) | 
| 47606 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 blanchet parents: 
47531diff
changeset | 174 | (#best_slices (config ()) ctxt) | 
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 175 | | NONE => false | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 176 | end | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 177 | |
| 44597 | 178 | val is_unit_equational_atp = is_atp_for_format (curry (op =) CNF_UEQ) | 
| 45303 
bd03b08161ac
added DFG unsorted support (like in the old days)
 blanchet parents: 
45301diff
changeset | 179 | val is_ho_atp = is_atp_for_format is_format_higher_order | 
| 44597 | 180 | |
| 45376 | 181 | fun is_prover_supported ctxt = | 
| 42361 | 182 | let val thy = Proof_Context.theory_of ctxt in | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 183 | is_reconstructor orf is_atp thy orf is_smt_prover ctxt | 
| 40941 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 blanchet parents: 
40723diff
changeset | 184 | end | 
| 40072 
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
 blanchet parents: 
40071diff
changeset | 185 | |
| 41432 
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
 boehmes parents: 
41336diff
changeset | 186 | fun is_prover_installed ctxt = | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 187 | is_reconstructor orf is_smt_prover ctxt orf | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 188 | is_atp_installed (Proof_Context.theory_of ctxt) | 
| 40941 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 blanchet parents: 
40723diff
changeset | 189 | |
| 45706 | 190 | fun get_slices slice slices = | 
| 191 | (0 upto length slices - 1) ~~ slices |> not slice ? (List.last #> single) | |
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 192 | |
| 48293 | 193 | val reconstructor_default_max_facts = 20 | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 194 | |
| 48293 | 195 | fun default_max_facts_for_prover ctxt slice name = | 
| 42361 | 196 | let val thy = Proof_Context.theory_of ctxt in | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 197 | if is_reconstructor name then | 
| 48293 | 198 | reconstructor_default_max_facts | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 199 | else if is_atp thy name then | 
| 46407 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46405diff
changeset | 200 | fold (Integer.max o #1 o fst o snd o snd o snd) | 
| 47606 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 blanchet parents: 
47531diff
changeset | 201 | (get_slices slice (#best_slices (get_atp thy name ()) ctxt)) 0 | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 202 | else (* is_smt_prover ctxt name *) | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 203 | SMT_Solver.default_max_relevant ctxt name | 
| 40941 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 blanchet parents: 
40723diff
changeset | 204 | end | 
| 40063 
d086e3699e78
bring ATPs and SMT solvers more in line with each other
 blanchet parents: 
40062diff
changeset | 205 | |
| 42956 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 206 | fun is_if (@{const_name If}, _) = true
 | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 207 | | is_if _ = false | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 208 | |
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 209 | (* Beware of "if and only if" (which is translated as such) and "If" (which is | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 210 | translated to conditional equations). *) | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 211 | fun is_good_unit_equality T t u = | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 212 |   T <> @{typ bool} andalso not (exists (exists_Const is_if) [t, u])
 | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 213 | |
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 214 | fun is_unit_equality (@{const Trueprop} $ t) = is_unit_equality t
 | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 215 |   | is_unit_equality (Const (@{const_name all}, _) $ Abs (_, _, t)) =
 | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 216 | is_unit_equality t | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 217 |   | is_unit_equality (Const (@{const_name All}, _) $ Abs (_, _, t)) =
 | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 218 | is_unit_equality t | 
| 42956 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 219 |   | is_unit_equality (Const (@{const_name "=="}, Type (_, [T, _])) $ t $ u) =
 | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 220 | is_good_unit_equality T t u | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 221 |   | is_unit_equality (Const (@{const_name HOL.eq}, Type (_ , [T, _])) $ t $ u) =
 | 
| 
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
 blanchet parents: 
42952diff
changeset | 222 | is_good_unit_equality T t u | 
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 223 | | is_unit_equality _ = false | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 224 | |
| 42952 
96f62b77748f
tuning -- the "appropriate" terminology is inspired from TPTP
 blanchet parents: 
42944diff
changeset | 225 | fun is_appropriate_prop_for_prover ctxt name = | 
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 226 | if is_unit_equational_atp ctxt name then is_unit_equality else K true | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 227 | |
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 228 | fun is_built_in_const_for_prover ctxt name = | 
| 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 229 | if is_smt_prover ctxt name then | 
| 41336 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 230 | let val ctxt = ctxt |> select_smt_solver name in | 
| 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 231 | fn x => fn ts => | 
| 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 232 | if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then | 
| 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 233 | (true, []) | 
| 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 234 | else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then | 
| 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 235 | (true, ts) | 
| 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 236 | else | 
| 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 237 | (false, ts) | 
| 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 238 | end | 
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 239 | else | 
| 41336 
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
 blanchet parents: 
41335diff
changeset | 240 | fn (s, _) => fn ts => (member (op =) atp_irrelevant_consts s, ts) | 
| 40071 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 241 | |
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 242 | (* FUDGE *) | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 243 | val atp_relevance_fudge = | 
| 42738 | 244 |   {local_const_multiplier = 1.5,
 | 
| 41159 
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
 blanchet parents: 
41152diff
changeset | 245 | worse_irrel_freq = 100.0, | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 246 | higher_order_irrel_weight = 1.05, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 247 | abs_rel_weight = 0.5, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 248 | abs_irrel_weight = 2.0, | 
| 47934 
08d7aff8c7e6
lower skolem penalty to ensure that some useful facts with existentials, e.g. congruence of "setsum", eventually get picked up
 blanchet parents: 
47912diff
changeset | 249 | skolem_irrel_weight = 0.05, | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 250 | theory_const_rel_weight = 0.5, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 251 | theory_const_irrel_weight = 0.25, | 
| 42735 
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
 blanchet parents: 
42730diff
changeset | 252 | chained_const_irrel_weight = 0.25, | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 253 | intro_bonus = 0.15, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 254 | elim_bonus = 0.15, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 255 | simp_bonus = 0.15, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 256 | local_bonus = 0.55, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 257 | assum_bonus = 1.05, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 258 | chained_bonus = 1.5, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 259 | max_imperfect = 11.5, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 260 | max_imperfect_exp = 1.0, | 
| 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 261 | threshold_divisor = 2.0, | 
| 41093 | 262 | ridiculous_threshold = 0.01} | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 263 | |
| 40071 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 264 | (* FUDGE (FIXME) *) | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 265 | val smt_relevance_fudge = | 
| 42738 | 266 |   {local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
 | 
| 41159 
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
 blanchet parents: 
41152diff
changeset | 267 | worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge, | 
| 40071 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 268 | higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 269 | abs_rel_weight = #abs_rel_weight atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 270 | abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 271 | skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 272 | theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 273 | theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge, | 
| 42735 
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
 blanchet parents: 
42730diff
changeset | 274 | chained_const_irrel_weight = #chained_const_irrel_weight atp_relevance_fudge, | 
| 40071 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 275 | intro_bonus = #intro_bonus atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 276 | elim_bonus = #elim_bonus atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 277 | simp_bonus = #simp_bonus atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 278 | local_bonus = #local_bonus atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 279 | assum_bonus = #assum_bonus atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 280 | chained_bonus = #chained_bonus atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 281 | max_imperfect = #max_imperfect atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 282 | max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 283 | threshold_divisor = #threshold_divisor atp_relevance_fudge, | 
| 
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
 blanchet parents: 
40070diff
changeset | 284 | ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge} | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 285 | |
| 40941 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 blanchet parents: 
40723diff
changeset | 286 | fun relevance_fudge_for_prover ctxt name = | 
| 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 blanchet parents: 
40723diff
changeset | 287 | if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge | 
| 40070 
bdb890782d4a
replaced references with proper record that's threaded through
 blanchet parents: 
40069diff
changeset | 288 | |
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41723diff
changeset | 289 | fun supported_provers ctxt = | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 290 | let | 
| 42361 | 291 | val thy = Proof_Context.theory_of ctxt | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 292 | val (remote_provers, local_provers) = | 
| 45520 | 293 | reconstructor_names @ | 
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41723diff
changeset | 294 | sort_strings (supported_atps thy) @ | 
| 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41723diff
changeset | 295 | sort_strings (SMT_Solver.available_solvers_of ctxt) | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 296 | |> List.partition (String.isPrefix remote_prefix) | 
| 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 297 | in | 
| 41727 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 blanchet parents: 
41723diff
changeset | 298 |     Output.urgent_message ("Supported provers: " ^
 | 
| 40205 
277508b07418
if "debug" is on, print list of relevant facts (poweruser request);
 blanchet parents: 
40204diff
changeset | 299 | commas (local_provers @ remote_provers) ^ ".") | 
| 40060 
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
 blanchet parents: 
40059diff
changeset | 300 | end | 
| 35969 | 301 | |
| 48319 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 blanchet parents: 
48314diff
changeset | 302 | fun kill_provers () = Async_Manager.kill_threads SledgehammerN "prover" | 
| 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 blanchet parents: 
48314diff
changeset | 303 | fun running_provers () = Async_Manager.running_threads SledgehammerN "prover" | 
| 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 blanchet parents: 
48314diff
changeset | 304 | val messages = Async_Manager.thread_messages SledgehammerN "prover" | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 305 | |
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 306 | |
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 307 | (** problems, results, ATPs, etc. **) | 
| 35969 | 308 | |
| 309 | type params = | |
| 48321 | 310 |   {debug : bool,
 | 
| 311 | verbose : bool, | |
| 312 | overlord : bool, | |
| 313 | blocking : bool, | |
| 314 | provers : string list, | |
| 315 | type_enc : string option, | |
| 316 | strict : bool, | |
| 317 | lam_trans : string option, | |
| 318 | uncurried_aliases : bool option, | |
| 319 | learn : bool, | |
| 320 | fact_filter : string option, | |
| 321 | max_facts : int option, | |
| 322 | fact_thresholds : real * real, | |
| 323 | max_mono_iters : int option, | |
| 324 | max_new_mono_instances : int option, | |
| 325 | isar_proof : bool, | |
| 326 | isar_shrink_factor : int, | |
| 327 | slice : bool, | |
| 328 | minimize : bool option, | |
| 329 | timeout : Time.time, | |
| 330 | preplay_timeout : Time.time, | |
| 331 | expect : string} | |
| 35867 | 332 | |
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 333 | type relevance_fudge = | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 334 |   {local_const_multiplier : real,
 | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 335 | worse_irrel_freq : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 336 | higher_order_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 337 | abs_rel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 338 | abs_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 339 | skolem_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 340 | theory_const_rel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 341 | theory_const_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 342 | chained_const_irrel_weight : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 343 | intro_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 344 | elim_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 345 | simp_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 346 | local_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 347 | assum_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 348 | chained_bonus : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 349 | max_imperfect : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 350 | max_imperfect_exp : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 351 | threshold_divisor : real, | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 352 | ridiculous_threshold : real} | 
| 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 353 | |
| 41090 | 354 | datatype prover_fact = | 
| 46340 | 355 | Untranslated_Fact of (string * stature) * thm | | 
| 356 | SMT_Weighted_Fact of (string * stature) * (int option * thm) | |
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 357 | |
| 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 358 | type prover_problem = | 
| 48321 | 359 |   {state : Proof.state,
 | 
| 360 | goal : thm, | |
| 361 | subgoal : int, | |
| 362 | subgoal_count : int, | |
| 363 | facts : prover_fact list} | |
| 35867 | 364 | |
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 365 | type prover_result = | 
| 48321 | 366 |   {outcome : failure option,
 | 
| 367 | used_facts : (string * stature) list, | |
| 368 | run_time : Time.time, | |
| 369 | preplay : unit -> play, | |
| 370 | message : play -> string, | |
| 371 | message_tail : string} | |
| 35867 | 372 | |
| 43051 | 373 | type prover = | 
| 45520 | 374 | params -> ((string * string list) list -> string -> minimize_command) | 
| 375 | -> prover_problem -> prover_result | |
| 35867 | 376 | |
| 38023 | 377 | (* configuration attributes *) | 
| 378 | ||
| 43092 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 blanchet parents: 
43085diff
changeset | 379 | (* Empty string means create files in Isabelle's temporary files directory. *) | 
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42593diff
changeset | 380 | val dest_dir = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42593diff
changeset | 381 |   Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
 | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42593diff
changeset | 382 | val problem_prefix = | 
| 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42593diff
changeset | 383 |   Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
 | 
| 48143 | 384 | val completish = | 
| 385 |   Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
 | |
| 28484 | 386 | |
| 43092 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 blanchet parents: 
43085diff
changeset | 387 | (* In addition to being easier to read, readable names are often much shorter, | 
| 44394 
20bd9f90accc
added option to control soundness of encodings more precisely, for evaluation purposes
 blanchet parents: 
44393diff
changeset | 388 | especially if types are mangled in names. This makes a difference for some | 
| 
20bd9f90accc
added option to control soundness of encodings more precisely, for evaluation purposes
 blanchet parents: 
44393diff
changeset | 389 | provers (e.g., E). For these reason, short names are enabled by default. *) | 
| 44592 
54906b0337ab
flip logic of boolean option so it's off by default
 blanchet parents: 
44586diff
changeset | 390 | val atp_full_names = | 
| 
54906b0337ab
flip logic of boolean option so it's off by default
 blanchet parents: 
44586diff
changeset | 391 |   Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
 | 
| 43092 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 blanchet parents: 
43085diff
changeset | 392 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 393 | val smt_triggers = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 394 |   Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 395 | val smt_weights = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 396 |   Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 397 | val smt_weight_min_facts = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 398 |   Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
 | 
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 399 | |
| 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 400 | (* FUDGE *) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 401 | val smt_min_weight = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 402 |   Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 403 | val smt_max_weight = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 404 |   Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 405 | val smt_max_weight_index = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 406 |   Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
 | 
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 407 | val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x) | 
| 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 408 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 409 | fun smt_fact_weight ctxt j num_facts = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 410 | if Config.get ctxt smt_weights andalso | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 411 | num_facts >= Config.get ctxt smt_weight_min_facts then | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 412 | let | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 413 | val min = Config.get ctxt smt_min_weight | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 414 | val max = Config.get ctxt smt_max_weight | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 415 | val max_index = Config.get ctxt smt_max_weight_index | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 416 | val curve = !smt_weight_curve | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 417 | in | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 418 | SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1)) | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 419 | div curve max_index) | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 420 | end | 
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 421 | else | 
| 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 422 | NONE | 
| 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 423 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 424 | fun weight_smt_fact ctxt num_facts ((info, th), j) = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 425 | let val thy = Proof_Context.theory_of ctxt in | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 426 | (info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy)) | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 427 | end | 
| 38023 | 428 | |
| 41091 
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
 blanchet parents: 
41090diff
changeset | 429 | fun untranslated_fact (Untranslated_Fact p) = p | 
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 430 | | untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th) | 
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 431 | fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 432 | | smt_weighted_fact ctxt num_facts (fact, j) = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 433 | (untranslated_fact fact, j) |> weight_smt_fact ctxt num_facts | 
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 434 | |
| 41313 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41259diff
changeset | 435 | fun overlord_file_location_for_prover prover = | 
| 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41259diff
changeset | 436 | (getenv "ISABELLE_HOME_USER", "prob_" ^ prover) | 
| 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41259diff
changeset | 437 | |
| 43033 | 438 | fun with_path cleanup after f path = | 
| 439 | Exn.capture f path | |
| 440 | |> tap (fn _ => cleanup path) | |
| 441 | |> Exn.release | |
| 442 | |> tap (after path) | |
| 443 | ||
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 444 | fun proof_banner mode name = | 
| 43033 | 445 | case mode of | 
| 446 |     Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof"
 | |
| 447 |   | Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
 | |
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 448 | | _ => "Try this" | 
| 43033 | 449 | |
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 450 | fun bunch_of_reconstructors needs_full_types lam_trans = | 
| 46405 | 451 | [(false, Metis (partial_type_enc, lam_trans false)), | 
| 46296 | 452 | (true, Metis (full_type_enc, lam_trans false)), | 
| 46405 | 453 | (false, Metis (no_typesN, lam_trans true)), | 
| 454 | (true, Metis (really_full_type_enc, lam_trans true)), | |
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 455 | (true, SMT)] | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 456 | |> map_filter (fn (full_types, reconstr) => | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 457 | if needs_full_types andalso not full_types then NONE | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 458 | else SOME reconstr) | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 459 | |
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 460 | fun extract_reconstructor ({type_enc, lam_trans, ...} : params)
 | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 461 | (Metis (type_enc', lam_trans')) = | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 462 | let | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 463 | val override_params = | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 464 | (if is_none type_enc andalso type_enc' = hd partial_type_encs then | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 465 | [] | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 466 | else | 
| 45566 
da05ce2de5a8
better threading of type encodings between Sledgehammer and "metis"
 blanchet parents: 
45561diff
changeset | 467 |            [("type_enc", [hd (unalias_type_enc type_enc')])]) @
 | 
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 468 | (if is_none lam_trans andalso lam_trans' = metis_default_lam_trans then | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 469 | [] | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 470 | else | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 471 |            [("lam_trans", [lam_trans'])])
 | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 472 | in (metisN, override_params) end | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 473 | | extract_reconstructor _ SMT = (smtN, []) | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 474 | |
| 43033 | 475 | (* based on "Mirabelle.can_apply" and generalized *) | 
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
43033diff
changeset | 476 | fun timed_apply timeout tac state i = | 
| 43033 | 477 | let | 
| 478 |     val {context = ctxt, facts, goal} = Proof.goal state
 | |
| 479 | val full_tac = Method.insert_tac facts i THEN tac ctxt i | |
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
43033diff
changeset | 480 | in TimeLimit.timeLimit timeout (try (Seq.pull o full_tac)) goal end | 
| 43033 | 481 | |
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 482 | fun tac_for_reconstructor (Metis (type_enc, lam_trans)) = | 
| 45521 | 483 | metis_tac [type_enc] lam_trans | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 484 | | tac_for_reconstructor SMT = SMT_Solver.smt_tac | 
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
43033diff
changeset | 485 | |
| 45520 | 486 | fun timed_reconstructor reconstr debug timeout ths = | 
| 44651 
5d6a11e166cf
renamed "Metis_Tactics" to "Metis_Tactic", now that there is only one Metis tactic ("metisFT" is legacy)
 blanchet parents: 
44649diff
changeset | 487 | (Config.put Metis_Tactic.verbose debug | 
| 45557 | 488 | #> Config.put SMT_Config.verbose debug | 
| 45520 | 489 | #> (fn ctxt => tac_for_reconstructor reconstr ctxt ths)) | 
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
43033diff
changeset | 490 | |> timed_apply timeout | 
| 43033 | 491 | |
| 492 | fun filter_used_facts used = filter (member (op =) used o fst) | |
| 493 | ||
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 494 | fun play_one_line_proof mode debug verbose timeout pairs state i preferred | 
| 45520 | 495 | reconstrs = | 
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
43033diff
changeset | 496 | let | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 497 | val _ = | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 498 | if mode = Minimize andalso Time.> (timeout, Time.zeroTime) then | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 499 | Output.urgent_message "Preplaying proof..." | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 500 | else | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 501 | () | 
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 502 | val ths = pairs |> sort_wrt (fst o fst) |> map snd | 
| 45520 | 503 | fun get_preferred reconstrs = | 
| 504 | if member (op =) reconstrs preferred then preferred | |
| 505 | else List.last reconstrs | |
| 506 | fun play [] [] = Failed_to_Play (get_preferred reconstrs) | |
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 507 | | play timed_outs [] = | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 508 | Trust_Playable (get_preferred timed_outs, SOME timeout) | 
| 45520 | 509 | | play timed_out (reconstr :: reconstrs) = | 
| 45378 | 510 | let | 
| 511 | val _ = | |
| 512 | if verbose then | |
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 513 | "Trying \"" ^ string_for_reconstructor reconstr ^ "\" for " ^ | 
| 45378 | 514 | string_from_time timeout ^ "..." | 
| 515 | |> Output.urgent_message | |
| 516 | else | |
| 517 | () | |
| 518 | val timer = Timer.startRealTimer () | |
| 519 | in | |
| 45520 | 520 | case timed_reconstructor reconstr debug timeout ths state i of | 
| 521 | SOME (SOME _) => Played (reconstr, Timer.checkRealTimer timer) | |
| 522 | | _ => play timed_out reconstrs | |
| 43044 
5945375700aa
always check plain "metis" even if the ATP proof seems to require "metisFT" -- maybe the proof is needlessly complicated
 blanchet parents: 
43037diff
changeset | 523 | end | 
| 45520 | 524 | handle TimeLimit.TimeOut => play (reconstr :: timed_out) reconstrs | 
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
43033diff
changeset | 525 | in | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 526 | if timeout = Time.zeroTime then | 
| 45520 | 527 | Trust_Playable (get_preferred reconstrs, NONE) | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 528 | else | 
| 45520 | 529 | play [] reconstrs | 
| 43033 | 530 | end | 
| 531 | ||
| 41313 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41259diff
changeset | 532 | |
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 533 | (* generic TPTP-based ATPs *) | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 534 | |
| 42730 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 535 | (* Too general means, positive equality literal with a variable X as one | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 536 | operand, when X does not occur properly in the other operand. This rules out | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 537 | clearly inconsistent facts such as X = a | X = b, though it by no means | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 538 | guarantees soundness. *) | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 539 | |
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 540 | (* Unwanted equalities are those between a (bound or schematic) variable that | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 541 | does not properly occur in the second operand. *) | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 542 | val is_exhaustive_finite = | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 543 | let | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 544 | fun is_bad_equal (Var z) t = | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 545 | not (exists_subterm (fn Var z' => z = z' | _ => false) t) | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 546 | | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j)) | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 547 | | is_bad_equal _ _ = false | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 548 | fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 549 | fun do_formula pos t = | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 550 | case (pos, t) of | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 551 |         (_, @{const Trueprop} $ t1) => do_formula pos t1
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 552 |       | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 553 | do_formula pos t' | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 554 |       | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 555 | do_formula pos t' | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 556 |       | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 557 | do_formula pos t' | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 558 |       | (_, @{const "==>"} $ t1 $ t2) =>
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 559 | do_formula (not pos) t1 andalso | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 560 |         (t2 = @{prop False} orelse do_formula pos t2)
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 561 |       | (_, @{const HOL.implies} $ t1 $ t2) =>
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 562 | do_formula (not pos) t1 andalso | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 563 |         (t2 = @{const False} orelse do_formula pos t2)
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 564 |       | (_, @{const Not} $ t1) => do_formula (not pos) t1
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 565 |       | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 566 |       | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 567 |       | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 568 |       | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
 | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 569 | | _ => false | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 570 | in do_formula true end | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 571 | |
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 572 | fun has_bound_or_var_of_type pred = | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 573 | exists_subterm (fn Var (_, T as Type _) => pred T | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 574 | | Abs (_, T as Type _, _) => pred T | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 575 | | _ => false) | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 576 | |
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 577 | (* Facts are forbidden to contain variables of these types. The typical reason | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 578 | is that they lead to unsoundness. Note that "unit" satisfies numerous | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 579 | equations like "?x = ()". The resulting clauses will have no type constraint, | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 580 | yielding false proofs. Even "bool" leads to many unsound proofs, though only | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 581 | for higher-order problems. *) | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 582 | |
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 583 | (* Facts containing variables of type "unit" or "bool" or of the form | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 584 | "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 585 | are omitted. *) | 
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 586 | fun is_dangerous_prop ctxt = | 
| 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42943diff
changeset | 587 | transform_elim_prop | 
| 44393 | 588 | #> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf | 
| 42730 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 589 | is_exhaustive_finite) | 
| 
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
 blanchet parents: 
42729diff
changeset | 590 | |
| 39492 | 591 | (* Important messages are important but not so important that users want to see | 
| 592 | them each time. *) | |
| 44649 | 593 | val atp_important_message_keep_quotient = 25 | 
| 39492 | 594 | |
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44397diff
changeset | 595 | fun choose_type_enc soundness best_type_enc format = | 
| 44397 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 blanchet parents: 
44394diff
changeset | 596 | the_default best_type_enc | 
| 
06375952f1fa
cleaner handling of polymorphic monotonicity inference
 blanchet parents: 
44394diff
changeset | 597 | #> type_enc_from_string soundness | 
| 44416 
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
 blanchet parents: 
44397diff
changeset | 598 | #> adjust_type_enc format | 
| 42548 
ea2a28b1938f
make sure the minimizer monomorphizes when it should
 blanchet parents: 
42544diff
changeset | 599 | |
| 43051 | 600 | val metis_minimize_max_time = seconds 2.0 | 
| 601 | ||
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 602 | fun choose_minimize_command params minimize_command name preplay = | 
| 45520 | 603 | let | 
| 604 | val (name, override_params) = | |
| 605 | case preplay of | |
| 606 | Played (reconstr, time) => | |
| 607 | if Time.<= (time, metis_minimize_max_time) then | |
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 608 | extract_reconstructor params reconstr | 
| 45520 | 609 | else | 
| 610 | (name, []) | |
| 611 | | _ => (name, []) | |
| 612 | in minimize_command override_params name end | |
| 43051 | 613 | |
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47946diff
changeset | 614 | fun repair_monomorph_context max_iters best_max_iters max_new_instances | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47946diff
changeset | 615 | best_max_new_instances = | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47946diff
changeset | 616 | Config.put Monomorph.max_rounds (max_iters |> the_default best_max_iters) | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47946diff
changeset | 617 | #> Config.put Monomorph.max_new_instances | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47946diff
changeset | 618 | (max_new_instances |> the_default best_max_new_instances) | 
| 43230 
dabf6e311213
clarified meaning of monomorphization configuration option by renaming it
 boehmes parents: 
43228diff
changeset | 619 | #> Config.put Monomorph.keep_partial_instances false | 
| 43226 | 620 | |
| 44509 
369e8c28a61a
added a component in generated file names reflecting whether the minimizer is used -- needed for evaluation to keep these files separated from the main problem files
 blanchet parents: 
44423diff
changeset | 621 | fun suffix_for_mode Auto_Try = "_auto_try" | 
| 
369e8c28a61a
added a component in generated file names reflecting whether the minimizer is used -- needed for evaluation to keep these files separated from the main problem files
 blanchet parents: 
44423diff
changeset | 622 | | suffix_for_mode Try = "_try" | 
| 
369e8c28a61a
added a component in generated file names reflecting whether the minimizer is used -- needed for evaluation to keep these files separated from the main problem files
 blanchet parents: 
44423diff
changeset | 623 | | suffix_for_mode Normal = "" | 
| 48392 
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
 blanchet parents: 
48376diff
changeset | 624 | | suffix_for_mode MaSh = "_mash" | 
| 45574 | 625 | | suffix_for_mode Auto_Minimize = "_auto_min" | 
| 44509 
369e8c28a61a
added a component in generated file names reflecting whether the minimizer is used -- needed for evaluation to keep these files separated from the main problem files
 blanchet parents: 
44423diff
changeset | 626 | | suffix_for_mode Minimize = "_min" | 
| 
369e8c28a61a
added a component in generated file names reflecting whether the minimizer is used -- needed for evaluation to keep these files separated from the main problem files
 blanchet parents: 
44423diff
changeset | 627 | |
| 44423 
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
 blanchet parents: 
44416diff
changeset | 628 | (* Give the ATPs some slack before interrupting them the hard way. "z3_tptp" on | 
| 43631 
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
 blanchet parents: 
43626diff
changeset | 629 | Linux appears to be the only ATP that does not honor its time limit. *) | 
| 43690 | 630 | val atp_timeout_slack = seconds 1.0 | 
| 43631 
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
 blanchet parents: 
43626diff
changeset | 631 | |
| 48331 
f190a6dbb29b
make the monomorphizer more predictable by making the cutoff independent on the number of facts
 blanchet parents: 
48321diff
changeset | 632 | val mono_max_privileged_facts = 10 | 
| 
f190a6dbb29b
make the monomorphizer more predictable by making the cutoff independent on the number of facts
 blanchet parents: 
48321diff
changeset | 633 | |
| 43021 | 634 | fun run_atp mode name | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 635 |         ({exec, arguments, proof_delims, known_failures, prem_role, best_slices,
 | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 636 | best_max_mono_iters, best_max_new_mono_instances, ...} : atp_config) | 
| 46301 | 637 |         (params as {debug, verbose, overlord, type_enc, strict, lam_trans,
 | 
| 48293 | 638 | uncurried_aliases, max_facts, max_mono_iters, | 
| 46409 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 639 | max_new_mono_instances, isar_proof, isar_shrink_factor, | 
| 
d4754183ccce
made option available to users (mostly for experiments)
 blanchet parents: 
46407diff
changeset | 640 | slice, timeout, preplay_timeout, ...}) | 
| 43037 | 641 | minimize_command | 
| 642 |         ({state, goal, subgoal, subgoal_count, facts, ...} : prover_problem) =
 | |
| 38023 | 643 | let | 
| 42182 
a630978fc967
start monomorphization process with subgoal, not entire goal, to avoid needless instances (and only print monomorphization messages in debug mode)
 blanchet parents: 
42181diff
changeset | 644 | val thy = Proof.theory_of state | 
| 39318 | 645 | val ctxt = Proof.context_of state | 
| 47946 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 blanchet parents: 
47934diff
changeset | 646 | val atp_mode = | 
| 48143 | 647 | if Config.get ctxt completish then Sledgehammer_Completish | 
| 47946 
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
 blanchet parents: 
47934diff
changeset | 648 | else Sledgehammer | 
| 43004 
20e9caff1f86
fix soundness bug in Sledgehammer: distinguish params in goals from fixed variables in context
 blanchet parents: 
42998diff
changeset | 649 | val (_, hyp_ts, concl_t) = strip_subgoal ctxt goal subgoal | 
| 41159 
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
 blanchet parents: 
41152diff
changeset | 650 | val (dest_dir, problem_prefix) = | 
| 
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
 blanchet parents: 
41152diff
changeset | 651 | if overlord then overlord_file_location_for_prover name | 
| 
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
 blanchet parents: 
41152diff
changeset | 652 | else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix) | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 653 | val problem_file_name = | 
| 41159 
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
 blanchet parents: 
41152diff
changeset | 654 | Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^ | 
| 44509 
369e8c28a61a
added a component in generated file names reflecting whether the minimizer is used -- needed for evaluation to keep these files separated from the main problem files
 blanchet parents: 
44423diff
changeset | 655 | suffix_for_mode mode ^ "_" ^ string_of_int subgoal) | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 656 | val problem_path_name = | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 657 | if dest_dir = "" then | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 658 | File.tmp_path problem_file_name | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 659 | else if File.exists (Path.explode dest_dir) then | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 660 | Path.append (Path.explode dest_dir) problem_file_name | 
| 39003 
c2aebd79981f
run relevance filter in a thread, to avoid blocking
 blanchet parents: 
39000diff
changeset | 661 | else | 
| 40059 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 blanchet parents: 
39494diff
changeset | 662 |         error ("No such directory: " ^ quote dest_dir ^ ".")
 | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 663 | val command0 = | 
| 47055 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47038diff
changeset | 664 | case find_first (fn var => getenv var <> "") (fst exec) of | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 665 | SOME var => | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 666 | let | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 667 | val pref = getenv var ^ "/" | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 668 | val paths = map (Path.explode o prefix pref) (snd exec) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 669 | in | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 670 | case find_first File.exists paths of | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 671 | SOME path => path | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 672 |           | NONE => error ("Bad executable: " ^ Path.print (hd paths) ^ ".")
 | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 673 | end | 
| 47055 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47038diff
changeset | 674 |       | NONE => error ("The environment variable " ^ quote (hd (fst exec)) ^
 | 
| 
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
 blanchet parents: 
47038diff
changeset | 675 | " is not set.") | 
| 38023 | 676 | fun split_time s = | 
| 677 | let | |
| 42448 | 678 | val split = String.tokens (fn c => str c = "\n") | 
| 47737 
63c939dcd055
made "split_last" more robust in the face of obscure low-level errors
 blanchet parents: 
47606diff
changeset | 679 | val (output, t) = | 
| 
63c939dcd055
made "split_last" more robust in the face of obscure low-level errors
 blanchet parents: 
47606diff
changeset | 680 | s |> split |> (try split_last #> the_default ([], "0")) | 
| 
63c939dcd055
made "split_last" more robust in the face of obscure low-level errors
 blanchet parents: 
47606diff
changeset | 681 | |>> cat_lines | 
| 42448 | 682 | fun as_num f = f >> (fst o read_int) | 
| 683 | val num = as_num (Scan.many1 Symbol.is_ascii_digit) | |
| 684 | val digit = Scan.one Symbol.is_ascii_digit | |
| 685 | val num3 = as_num (digit ::: digit ::: (digit >> single)) | |
| 686 | val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b) | |
| 45381 | 687 | val as_time = | 
| 688 | raw_explode #> Scan.read Symbol.stopper time #> the_default 0 | |
| 47737 
63c939dcd055
made "split_last" more robust in the face of obscure low-level errors
 blanchet parents: 
47606diff
changeset | 689 | in (output, as_time t |> Time.fromMilliseconds) end | 
| 41313 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41259diff
changeset | 690 | fun run_on prob_file = | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 691 | let | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 692 | (* If slicing is disabled, we expand the last slice to fill the entire | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 693 | time available. *) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 694 | val actual_slices = get_slices slice (best_slices ctxt) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 695 | val num_actual_slices = length actual_slices | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 696 | fun monomorphize_facts facts = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 697 | let | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 698 | val ctxt = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 699 | ctxt | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 700 | |> repair_monomorph_context max_mono_iters best_max_mono_iters | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 701 | max_new_mono_instances best_max_new_mono_instances | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 702 | (* pseudo-theorem involving the same constants as the subgoal *) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 703 | val subgoal_th = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 704 | Logic.list_implies (hyp_ts, concl_t) |> Skip_Proof.make_thm thy | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 705 | val rths = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 706 | facts |> chop mono_max_privileged_facts | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 707 | |>> map (pair 1 o snd) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 708 | ||> map (pair 2 o snd) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 709 | |> op @ | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 710 | |> cons (0, subgoal_th) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 711 | in | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 712 | Monomorph.monomorph atp_schematic_consts_of rths ctxt |> fst |> tl | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 713 | |> curry ListPair.zip (map fst facts) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 714 | |> maps (fn (name, rths) => | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 715 | map (pair name o zero_var_indexes o snd) rths) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 716 | end | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 717 | fun run_slice time_left (cache_key, cache_value) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 718 | (slice, (time_frac, (complete, | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 719 | (key as (best_max_facts, format, best_type_enc, | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 720 | best_lam_trans, best_uncurried_aliases), | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 721 | extra)))) = | 
| 38032 | 722 | let | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 723 | val num_facts = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 724 | length facts |> is_none max_facts ? Integer.min best_max_facts | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 725 | val soundness = if strict then Strict else Non_Strict | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 726 | val type_enc = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 727 | type_enc |> choose_type_enc soundness best_type_enc format | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 728 | val sound = is_type_enc_sound type_enc | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 729 | val real_ms = Real.fromInt o Time.toMilliseconds | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 730 | val slice_timeout = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 731 | ((real_ms time_left | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 732 | |> (if slice < num_actual_slices - 1 then | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 733 | curry Real.min (time_frac * real_ms timeout) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 734 | else | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 735 | I)) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 736 | * 0.001) |> seconds | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 737 | val generous_slice_timeout = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 738 | Time.+ (slice_timeout, atp_timeout_slack) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 739 | val _ = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 740 | if debug then | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 741 | quote name ^ " slice #" ^ string_of_int (slice + 1) ^ | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 742 | " with " ^ string_of_int num_facts ^ " fact" ^ | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 743 | plural_s num_facts ^ " for " ^ | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 744 | string_from_time slice_timeout ^ "..." | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 745 | |> Output.urgent_message | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 746 | else | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 747 | () | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 748 | val readable_names = not (Config.get ctxt atp_full_names) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 749 | val lam_trans = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 750 | case lam_trans of | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 751 | SOME s => s | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 752 | | NONE => best_lam_trans | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 753 | val uncurried_aliases = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 754 | case uncurried_aliases of | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 755 | SOME b => b | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 756 | | NONE => best_uncurried_aliases | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 757 | val value as (atp_problem, _, fact_names, _, _) = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 758 | if cache_key = SOME key then | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 759 | cache_value | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 760 | else | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 761 | facts | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 762 | |> map untranslated_fact | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 763 | |> not sound | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 764 | ? filter_out (is_dangerous_prop ctxt o prop_of o snd) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 765 | |> take num_facts | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 766 | |> not (is_type_enc_polymorphic type_enc) ? monomorphize_facts | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 767 | |> map (apsnd prop_of) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 768 | |> prepare_atp_problem ctxt format prem_role type_enc atp_mode | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 769 | lam_trans uncurried_aliases | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 770 | readable_names true hyp_ts concl_t | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 771 | fun sel_weights () = atp_problem_selection_weights atp_problem | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 772 | fun ord_info () = atp_problem_term_order_info atp_problem | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 773 | val ord = effective_term_order ctxt name | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 774 | val full_proof = debug orelse isar_proof | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 775 | val args = arguments ctxt full_proof extra slice_timeout | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 776 | (ord, ord_info, sel_weights) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 777 | val command = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 778 | File.shell_path command0 ^ " " ^ args ^ " " ^ | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 779 | File.shell_path prob_file | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 780 |               |> enclose "TIMEFORMAT='%3R'; { time " " ; } 2>&1"
 | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 781 | val _ = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 782 | atp_problem | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 783 | |> lines_for_atp_problem format ord ord_info | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 784 |               |> cons ("% " ^ command ^ "\n")
 | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 785 | |> File.write_list prob_file | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 786 | val ((output, run_time), (atp_proof, outcome)) = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 787 | TimeLimit.timeLimit generous_slice_timeout | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 788 | Isabelle_System.bash_output command | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 789 | |>> (if overlord then | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 790 |                      prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
 | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 791 | else | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 792 | I) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 793 | |> fst |> split_time | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 794 | |> (fn accum as (output, _) => | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 795 | (accum, | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 796 | extract_tstplike_proof_and_outcome verbose complete | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 797 | proof_delims known_failures output | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 798 | |>> atp_proof_from_tstplike_proof atp_problem | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 799 | handle UNRECOGNIZED_ATP_PROOF () => | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 800 | ([], SOME ProofIncomplete))) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 801 | handle TimeLimit.TimeOut => | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 802 |                      (("", slice_timeout), ([], SOME TimedOut))
 | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 803 | val outcome = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 804 | case outcome of | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 805 | NONE => | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 806 | (case used_facts_in_unsound_atp_proof ctxt fact_names atp_proof | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 807 | |> Option.map (sort string_ord) of | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 808 | SOME facts => | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 809 | let | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 810 | val failure = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 811 | UnsoundProof (is_type_enc_sound type_enc, facts) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 812 | in | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 813 | if debug then (warning (string_for_failure failure); NONE) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 814 | else SOME failure | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 815 | end | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 816 | | NONE => NONE) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 817 | | _ => outcome | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 818 | in ((SOME key, value), (output, run_time, atp_proof, outcome)) end | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 819 | val timer = Timer.startRealTimer () | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 820 | fun maybe_run_slice slice | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 821 | (result as (cache, (_, run_time0, _, SOME _))) = | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 822 | let | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 823 | val time_left = Time.- (timeout, Timer.checkRealTimer timer) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 824 | in | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 825 | if Time.<= (time_left, Time.zeroTime) then | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 826 | result | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 827 | else | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 828 | run_slice time_left cache slice | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 829 | |> (fn (cache, (output, run_time, atp_proof, outcome)) => | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 830 | (cache, (output, Time.+ (run_time0, run_time), | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 831 | atp_proof, outcome))) | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 832 | end | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 833 | | maybe_run_slice _ result = result | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 834 | in | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 835 | ((NONE, ([], Symtab.empty, Vector.fromList [], [], Symtab.empty)), | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 836 |          ("", Time.zeroTime, [], SOME InternalError))
 | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 837 | |> fold maybe_run_slice actual_slices | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 838 | end | 
| 38023 | 839 | |
| 840 | (* If the problem file has not been exported, remove it; otherwise, export | |
| 841 | the proof file too. *) | |
| 41313 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41259diff
changeset | 842 | fun cleanup prob_file = | 
| 
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
 blanchet parents: 
41259diff
changeset | 843 | if dest_dir = "" then try File.rm prob_file else NONE | 
| 43304 
6901ebafbb8d
cleaner fact freshening, which also works in corner cases, e.g. if two backquoted facts have the same name (but have different variable indices)
 blanchet parents: 
43303diff
changeset | 844 | fun export prob_file (_, (output, _, _, _)) = | 
| 48376 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 845 | if dest_dir = "" then () | 
| 
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
 blanchet parents: 
48331diff
changeset | 846 | else File.write (Path.explode (Path.implode prob_file ^ "_proof")) output | 
| 46407 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46405diff
changeset | 847 | val ((_, (_, pool, fact_names, _, sym_tab)), | 
| 
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
 blanchet parents: 
46405diff
changeset | 848 | (output, run_time, atp_proof, outcome)) = | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 849 | with_path cleanup export run_on problem_path_name | 
| 39492 | 850 | val important_message = | 
| 43021 | 851 | if mode = Normal andalso | 
| 42609 | 852 | random_range 0 (atp_important_message_keep_quotient - 1) = 0 then | 
| 39492 | 853 | extract_important_message output | 
| 854 | else | |
| 855 | "" | |
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 856 | val (used_facts, preplay, message, message_tail) = | 
| 38023 | 857 | case outcome of | 
| 858 | NONE => | |
| 43033 | 859 | let | 
| 45551 | 860 | val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof | 
| 45590 | 861 | val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof | 
| 45521 | 862 | val reconstrs = | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 863 | bunch_of_reconstructors needs_full_types | 
| 45560 
1606122a2d0f
don't needlessly pass "lam_lifted" option to "metis" call for SMT proof
 blanchet parents: 
45557diff
changeset | 864 | (lam_trans_from_atp_proof atp_proof | 
| 46405 | 865 | o (fn desperate => if desperate then hide_lamsN | 
| 866 | else metis_default_lam_trans)) | |
| 43033 | 867 | in | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 868 | (used_facts, | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 869 | fn () => | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 870 | let | 
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 871 | val used_pairs = | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 872 | facts |> map untranslated_fact |> filter_used_facts used_facts | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 873 | in | 
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 874 | play_one_line_proof mode debug verbose preplay_timeout | 
| 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 875 | used_pairs state subgoal (hd reconstrs) reconstrs | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 876 | end, | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 877 | fn preplay => | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 878 | let | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 879 | val isar_params = | 
| 45552 | 880 | (debug, isar_shrink_factor, pool, fact_names, sym_tab, | 
| 881 | atp_proof, goal) | |
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 882 | val one_line_params = | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 883 | (preplay, proof_banner mode name, used_facts, | 
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 884 | choose_minimize_command params minimize_command name preplay, | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 885 | subgoal, subgoal_count) | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 886 | in proof_text ctxt isar_proof isar_params one_line_params end, | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 887 | (if verbose then | 
| 45381 | 888 | "\nATP real CPU time: " ^ string_from_time run_time ^ "." | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 889 | else | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 890 | "") ^ | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 891 | (if important_message <> "" then | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 892 | "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 893 | important_message | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 894 | else | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 895 | "")) | 
| 43033 | 896 | end | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 897 | | SOME failure => | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 898 | ([], K (Failed_to_Play plain_metis), | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 899 | fn _ => string_for_failure failure, "") | 
| 38023 | 900 | in | 
| 45381 | 901 |     {outcome = outcome, used_facts = used_facts, run_time = run_time,
 | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 902 | preplay = preplay, message = message, message_tail = message_tail} | 
| 38023 | 903 | end | 
| 904 | ||
| 40669 | 905 | (* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until | 
| 906 | these are sorted out properly in the SMT module, we have to interpret these | |
| 907 | ourselves. *) | |
| 40684 
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
 blanchet parents: 
40669diff
changeset | 908 | val remote_smt_failures = | 
| 43631 
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
 blanchet parents: 
43626diff
changeset | 909 | [(2, NoLibwwwPerl), | 
| 
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
 blanchet parents: 
43626diff
changeset | 910 | (22, CantConnect)] | 
| 40684 
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
 blanchet parents: 
40669diff
changeset | 911 | val z3_failures = | 
| 41236 | 912 | [(101, OutOfResources), | 
| 913 | (103, MalformedInput), | |
| 41222 | 914 | (110, MalformedInput)] | 
| 40684 
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
 blanchet parents: 
40669diff
changeset | 915 | val unix_failures = | 
| 
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
 blanchet parents: 
40669diff
changeset | 916 | [(139, Crashed)] | 
| 43631 
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
 blanchet parents: 
43626diff
changeset | 917 | val smt_failures = remote_smt_failures @ z3_failures @ unix_failures | 
| 40555 | 918 | |
| 42100 
062381c5f9f8
more precise failure reporting in Sledgehammer/SMT
 blanchet parents: 
42061diff
changeset | 919 | fun failure_from_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
 | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 920 | if is_real_cex then Unprovable else GaveUp | 
| 41222 | 921 | | failure_from_smt_failure SMT_Failure.Time_Out = TimedOut | 
| 922 | | failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) = | |
| 923 | (case AList.lookup (op =) smt_failures code of | |
| 40684 
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
 blanchet parents: 
40669diff
changeset | 924 | SOME failure => failure | 
| 41259 
13972ced98d9
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
 blanchet parents: 
41256diff
changeset | 925 |      | NONE => UnknownError ("Abnormal termination with exit code " ^
 | 
| 
13972ced98d9
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
 blanchet parents: 
41256diff
changeset | 926 | string_of_int code ^ ".")) | 
| 41222 | 927 | | failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources | 
| 928 | | failure_from_smt_failure (SMT_Failure.Other_Failure msg) = | |
| 42061 
71077681eaf6
let SMT errors through -- the main reason for keeping them quiet was that the SMT bridge used to suffer from internal bugs, but these have been fixed for some time now
 blanchet parents: 
42060diff
changeset | 929 | UnknownError msg | 
| 40063 
d086e3699e78
bring ATPs and SMT solvers more in line with each other
 blanchet parents: 
40062diff
changeset | 930 | |
| 40698 | 931 | (* FUDGE *) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 932 | val smt_max_slices = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 933 |   Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 934 | val smt_slice_fact_frac = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 935 |   Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.5)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 936 | val smt_slice_time_frac = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 937 |   Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.5)
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 938 | val smt_slice_min_secs = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 939 |   Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 5)
 | 
| 40409 
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
 blanchet parents: 
40370diff
changeset | 940 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 941 | fun smt_filter_loop ctxt name | 
| 42724 
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
 blanchet parents: 
42723diff
changeset | 942 |                     ({debug, verbose, overlord, max_mono_iters,
 | 
| 45706 | 943 | max_new_mono_instances, timeout, slice, ...} : params) | 
| 47531 
7fe7c7419489
get rid of minor optimization that caused strange problems and was hard to debug (and apparently saved less than 100 ms on a 30 s run)
 blanchet parents: 
47055diff
changeset | 944 | state i = | 
| 40409 
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
 blanchet parents: 
40370diff
changeset | 945 | let | 
| 45706 | 946 | val max_slices = if slice then Config.get ctxt smt_max_slices else 1 | 
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 947 | val repair_context = | 
| 43233 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 948 | select_smt_solver name | 
| 45557 | 949 | #> Config.put SMT_Config.verbose debug | 
| 43233 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 950 | #> (if overlord then | 
| 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 951 | Config.put SMT_Config.debug_files | 
| 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 952 | (overlord_file_location_for_prover name | 
| 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 953 | |> (fn (path, name) => path ^ "/" ^ name)) | 
| 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 954 | else | 
| 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 955 | I) | 
| 
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
 blanchet parents: 
43232diff
changeset | 956 | #> Config.put SMT_Config.infer_triggers (Config.get ctxt smt_triggers) | 
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 957 | val state = state |> Proof.map_context repair_context | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 958 | fun do_slice timeout slice outcome0 time_so_far facts = | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 959 | let | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 960 | val timer = Timer.startRealTimer () | 
| 42724 
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
 blanchet parents: 
42723diff
changeset | 961 | val state = | 
| 
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
 blanchet parents: 
42723diff
changeset | 962 | state |> Proof.map_context | 
| 43267 | 963 | (repair_monomorph_context max_mono_iters | 
| 47962 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47946diff
changeset | 964 | default_max_mono_iters max_new_mono_instances | 
| 
137883567114
lower the monomorphization thresholds for less scalable provers
 blanchet parents: 
47946diff
changeset | 965 | default_max_new_mono_instances) | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 966 | val ms = timeout |> Time.toMilliseconds | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 967 | val slice_timeout = | 
| 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 968 | if slice < max_slices then | 
| 41169 | 969 | Int.min (ms, | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 970 | Int.max (1000 * Config.get ctxt smt_slice_min_secs, | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 971 | Real.ceil (Config.get ctxt smt_slice_time_frac | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 972 | * Real.fromInt ms))) | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 973 | |> Time.fromMilliseconds | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 974 | else | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 975 | timeout | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 976 | val num_facts = length facts | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 977 | val _ = | 
| 42614 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 978 | if debug then | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 979 | quote name ^ " slice " ^ string_of_int slice ^ " with " ^ | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 980 | string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ " for " ^ | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 981 | string_from_time slice_timeout ^ "..." | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 982 | |> Output.urgent_message | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 983 | else | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 984 | () | 
| 41168 | 985 | val birth = Timer.checkRealTimer timer | 
| 41171 | 986 | val _ = | 
| 41211 
1e2e16bc0077
no need to do a super-duper atomization if Metis fails afterwards anyway
 blanchet parents: 
41209diff
changeset | 987 | if debug then Output.urgent_message "Invoking SMT solver..." else () | 
| 41209 | 988 | val (outcome, used_facts) = | 
| 47531 
7fe7c7419489
get rid of minor optimization that caused strange problems and was hard to debug (and apparently saved less than 100 ms on a 30 s run)
 blanchet parents: 
47055diff
changeset | 989 | SMT_Solver.smt_filter_preprocess state facts i | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 990 | |> SMT_Solver.smt_filter_apply slice_timeout | 
| 41239 | 991 |           |> (fn {outcome, used_facts} => (outcome, used_facts))
 | 
| 41209 | 992 | handle exn => if Exn.is_interrupt exn then | 
| 993 | reraise exn | |
| 994 | else | |
| 42061 
71077681eaf6
let SMT errors through -- the main reason for keeping them quiet was that the SMT bridge used to suffer from internal bugs, but these have been fixed for some time now
 blanchet parents: 
42060diff
changeset | 995 | (ML_Compiler.exn_message exn | 
| 41209 | 996 | |> SMT_Failure.Other_Failure |> SOME, []) | 
| 41168 | 997 | val death = Timer.checkRealTimer timer | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 998 | val outcome0 = if is_none outcome0 then SOME outcome else outcome0 | 
| 41168 | 999 | val time_so_far = Time.+ (time_so_far, Time.- (death, birth)) | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1000 | val too_many_facts_perhaps = | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1001 | case outcome of | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1002 | NONE => false | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1003 | | SOME (SMT_Failure.Counterexample _) => false | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 1004 | | SOME SMT_Failure.Time_Out => slice_timeout <> timeout | 
| 42614 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1005 | | SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *) | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1006 | | SOME SMT_Failure.Out_Of_Memory => true | 
| 41211 
1e2e16bc0077
no need to do a super-duper atomization if Metis fails afterwards anyway
 blanchet parents: 
41209diff
changeset | 1007 | | SOME (SMT_Failure.Other_Failure _) => true | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1008 | val timeout = Time.- (timeout, Timer.checkRealTimer timer) | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1009 | in | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 1010 | if too_many_facts_perhaps andalso slice < max_slices andalso | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1011 | num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then | 
| 41169 | 1012 | let | 
| 42614 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1013 | val new_num_facts = | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 1014 | Real.ceil (Config.get ctxt smt_slice_fact_frac | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 1015 | * Real.fromInt num_facts) | 
| 42614 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1016 | val _ = | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1017 | if verbose andalso is_some outcome then | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1018 | quote name ^ " invoked with " ^ string_of_int num_facts ^ | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1019 | " fact" ^ plural_s num_facts ^ ": " ^ | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1020 | string_for_failure (failure_from_smt_failure (the outcome)) ^ | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1021 | " Retrying with " ^ string_of_int new_num_facts ^ " fact" ^ | 
| 42638 
a7a30721767a
have each ATP filter out dangerous facts for themselves, based on their type system
 blanchet parents: 
42623diff
changeset | 1022 | plural_s new_num_facts ^ "..." | 
| 42614 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1023 | |> Output.urgent_message | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1024 | else | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1025 | () | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 1026 | in | 
| 42614 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1027 | facts |> take new_num_facts | 
| 
81953e554197
make "debug" more verbose and "verbose" less verbose
 blanchet parents: 
42613diff
changeset | 1028 | |> do_slice timeout (slice + 1) outcome0 time_so_far | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 1029 | end | 
| 40553 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1030 | else | 
| 
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
 blanchet parents: 
40471diff
changeset | 1031 |           {outcome = if is_none outcome then NONE else the outcome0,
 | 
| 45370 
bab52dafa63a
use "Time.time" rather than milliseconds internally
 blanchet parents: 
45369diff
changeset | 1032 | used_facts = used_facts, run_time = time_so_far} | 
| 40409 
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
 blanchet parents: 
40370diff
changeset | 1033 | end | 
| 42443 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 blanchet parents: 
42361diff
changeset | 1034 | in do_slice timeout 1 NONE Time.zeroTime end | 
| 40409 
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
 blanchet parents: 
40370diff
changeset | 1035 | |
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1036 | fun run_smt_solver mode name (params as {debug, verbose, preplay_timeout, ...})
 | 
| 43011 
5f8d74d3b297
added syntax for specifying Metis timeout (currently used only by SMT solvers)
 blanchet parents: 
43006diff
changeset | 1037 | minimize_command | 
| 47531 
7fe7c7419489
get rid of minor optimization that caused strange problems and was hard to debug (and apparently saved less than 100 ms on a 30 s run)
 blanchet parents: 
47055diff
changeset | 1038 |         ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
 | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 1039 | let | 
| 41242 
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
 blanchet parents: 
41241diff
changeset | 1040 | val ctxt = Proof.context_of state | 
| 41256 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 1041 | val num_facts = length facts | 
| 
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
 blanchet parents: 
41242diff
changeset | 1042 | val facts = facts ~~ (0 upto num_facts - 1) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42642diff
changeset | 1043 | |> map (smt_weighted_fact ctxt num_facts) | 
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 1044 |     val {outcome, used_facts = used_pairs, run_time} =
 | 
| 47531 
7fe7c7419489
get rid of minor optimization that caused strange problems and was hard to debug (and apparently saved less than 100 ms on a 30 s run)
 blanchet parents: 
47055diff
changeset | 1045 | smt_filter_loop ctxt name params state subgoal facts | 
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 1046 | val used_facts = used_pairs |> map fst | 
| 41222 | 1047 | val outcome = outcome |> Option.map failure_from_smt_failure | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 1048 | val (preplay, message, message_tail) = | 
| 40184 
91b4b73dbafb
proper error handling for SMT solvers in Sledgehammer
 blanchet parents: 
40181diff
changeset | 1049 | case outcome of | 
| 
91b4b73dbafb
proper error handling for SMT solvers in Sledgehammer
 blanchet parents: 
40181diff
changeset | 1050 | NONE => | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1051 | (fn () => | 
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 1052 | play_one_line_proof mode debug verbose preplay_timeout used_pairs | 
| 45554 
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
 blanchet parents: 
45553diff
changeset | 1053 | state subgoal SMT | 
| 45560 
1606122a2d0f
don't needlessly pass "lam_lifted" option to "metis" call for SMT proof
 blanchet parents: 
45557diff
changeset | 1054 | (bunch_of_reconstructors false | 
| 
1606122a2d0f
don't needlessly pass "lam_lifted" option to "metis" call for SMT proof
 blanchet parents: 
45557diff
changeset | 1055 | (fn plain => | 
| 46365 | 1056 | if plain then metis_default_lam_trans else liftingN)), | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1057 | fn preplay => | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1058 | let | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1059 | val one_line_params = | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1060 | (preplay, proof_banner mode name, used_facts, | 
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1061 | choose_minimize_command params minimize_command name preplay, | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1062 | subgoal, subgoal_count) | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 1063 | in one_line_proof_text one_line_params end, | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 1064 | if verbose then | 
| 45370 
bab52dafa63a
use "Time.time" rather than milliseconds internally
 blanchet parents: 
45369diff
changeset | 1065 | "\nSMT solver real CPU time: " ^ string_from_time run_time ^ "." | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 1066 | else | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 1067 | "") | 
| 43166 | 1068 | | SOME failure => | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 1069 | (K (Failed_to_Play plain_metis), fn _ => string_for_failure failure, "") | 
| 40063 
d086e3699e78
bring ATPs and SMT solvers more in line with each other
 blanchet parents: 
40062diff
changeset | 1070 | in | 
| 45370 
bab52dafa63a
use "Time.time" rather than milliseconds internally
 blanchet parents: 
45369diff
changeset | 1071 |     {outcome = outcome, used_facts = used_facts, run_time = run_time,
 | 
| 
bab52dafa63a
use "Time.time" rather than milliseconds internally
 blanchet parents: 
45369diff
changeset | 1072 | preplay = preplay, message = message, message_tail = message_tail} | 
| 40063 
d086e3699e78
bring ATPs and SMT solvers more in line with each other
 blanchet parents: 
40062diff
changeset | 1073 | end | 
| 
d086e3699e78
bring ATPs and SMT solvers more in line with each other
 blanchet parents: 
40062diff
changeset | 1074 | |
| 45520 | 1075 | fun run_reconstructor mode name | 
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1076 |         (params as {debug, verbose, timeout, type_enc, lam_trans, ...})
 | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 1077 | minimize_command | 
| 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 1078 |         ({state, subgoal, subgoal_count, facts, ...} : prover_problem) =
 | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 1079 | let | 
| 45520 | 1080 | val reconstr = | 
| 1081 | if name = metisN then | |
| 1082 | Metis (type_enc |> the_default (hd partial_type_encs), | |
| 1083 | lam_trans |> the_default metis_default_lam_trans) | |
| 1084 | else if name = smtN then | |
| 1085 | SMT | |
| 1086 | else | |
| 1087 |         raise Fail ("unknown reconstructor: " ^ quote name)
 | |
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 1088 | val used_pairs = facts |> map untranslated_fact | 
| 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 1089 | val used_facts = used_pairs |> map fst | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 1090 | in | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 1091 | case play_one_line_proof (if mode = Minimize then Normal else mode) debug | 
| 45781 
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
 blanchet parents: 
45707diff
changeset | 1092 | verbose timeout used_pairs state subgoal reconstr | 
| 45520 | 1093 | [reconstr] of | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 1094 | play as Played (_, time) => | 
| 45370 
bab52dafa63a
use "Time.time" rather than milliseconds internally
 blanchet parents: 
45369diff
changeset | 1095 |       {outcome = NONE, used_facts = used_facts, run_time = time,
 | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1096 | preplay = K play, | 
| 45561 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1097 | message = | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1098 | fn play => | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1099 | let | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1100 | val (_, override_params) = extract_reconstructor params reconstr | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1101 | val one_line_params = | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1102 | (play, proof_banner mode name, used_facts, | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1103 | minimize_command override_params name, subgoal, | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1104 | subgoal_count) | 
| 
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
 blanchet parents: 
45560diff
changeset | 1105 | in one_line_proof_text one_line_params end, | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 1106 | message_tail = ""} | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1107 | | play => | 
| 43166 | 1108 | let | 
| 1109 | val failure = case play of Failed_to_Play _ => GaveUp | _ => TimedOut | |
| 1110 | in | |
| 45370 
bab52dafa63a
use "Time.time" rather than milliseconds internally
 blanchet parents: 
45369diff
changeset | 1111 |         {outcome = SOME failure, used_facts = [], run_time = Time.zeroTime,
 | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 1112 | preplay = K play, message = fn _ => string_for_failure failure, | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 1113 | message_tail = ""} | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 1114 | end | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 1115 | end | 
| 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43044diff
changeset | 1116 | |
| 43021 | 1117 | fun get_prover ctxt mode name = | 
| 42361 | 1118 | let val thy = Proof_Context.theory_of ctxt in | 
| 45379 
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
 blanchet parents: 
45378diff
changeset | 1119 | if is_reconstructor name then run_reconstructor mode name | 
| 47606 
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
 blanchet parents: 
47531diff
changeset | 1120 | else if is_atp thy name then run_atp mode name (get_atp thy name ()) | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1121 | else if is_smt_prover ctxt name then run_smt_solver mode name | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 1122 |     else error ("No such prover: " ^ name ^ ".")
 | 
| 40941 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 blanchet parents: 
40723diff
changeset | 1123 | end | 
| 40063 
d086e3699e78
bring ATPs and SMT solvers more in line with each other
 blanchet parents: 
40062diff
changeset | 1124 | |
| 28582 | 1125 | end; |