author | blanchet |
Tue, 31 May 2011 16:38:36 +0200 | |
changeset 43092 | 93ec303e1917 |
parent 43085 | 0a2f5b86bdd7 |
child 43096 | f181d66046d4 |
permissions | -rw-r--r-- |
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
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:
32995
diff
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:
41066
diff
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:
41066
diff
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 |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
12 |
type locality = ATP_Translate.locality |
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
13 |
type relevance_fudge = Sledgehammer_Filter.relevance_fudge |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
14 |
type translated_formula = ATP_Translate.translated_formula |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
15 |
type type_system = ATP_Translate.type_system |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
16 |
type play = ATP_Reconstruct.play |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
17 |
type minimize_command = ATP_Reconstruct.minimize_command |
39493
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents:
39492
diff
changeset
|
18 |
|
43021 | 19 |
datatype mode = Auto_Try | Try | Normal | Minimize |
20 |
||
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
21 |
datatype rich_type_system = |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
22 |
Dumb_Type_Sys of type_system | |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
23 |
Smart_Type_Sys of bool |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
24 |
|
35969 | 25 |
type params = |
41208
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
41171
diff
changeset
|
26 |
{debug: bool, |
35969 | 27 |
verbose: bool, |
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36064
diff
changeset
|
28 |
overlord: bool, |
41208
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
41171
diff
changeset
|
29 |
blocking: bool, |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
30 |
provers: string list, |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
31 |
type_sys: rich_type_system, |
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
42100
diff
changeset
|
32 |
relevance_thresholds: real * real, |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
42100
diff
changeset
|
33 |
max_relevant: int option, |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
34 |
max_mono_iters: int, |
42740
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents:
42738
diff
changeset
|
35 |
max_new_mono_instances: int, |
43063 | 36 |
explicit_apply: bool option, |
35969 | 37 |
isar_proof: bool, |
36924 | 38 |
isar_shrink_factor: int, |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
39 |
slicing: bool, |
38985
162bbbea4e4d
added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents:
38982
diff
changeset
|
40 |
timeout: Time.time, |
43015
21b6baec55b1
renamed "metis_timeout" to "preplay_timeout" and continued implementation
blanchet
parents:
43011
diff
changeset
|
41 |
preplay_timeout: Time.time, |
38985
162bbbea4e4d
added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents:
38982
diff
changeset
|
42 |
expect: string} |
39493
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents:
39492
diff
changeset
|
43 |
|
41090 | 44 |
datatype prover_fact = |
45 |
Untranslated_Fact of (string * locality) * thm | |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
46 |
SMT_Weighted_Fact of (string * locality) * (int option * thm) |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
47 |
|
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
48 |
type prover_problem = |
39318 | 49 |
{state: Proof.state, |
38998 | 50 |
goal: thm, |
51 |
subgoal: int, |
|
40065 | 52 |
subgoal_count: int, |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
53 |
facts: prover_fact list, |
41741 | 54 |
smt_filter: (string * locality) SMT_Solver.smt_filter_data option} |
39493
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents:
39492
diff
changeset
|
55 |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
56 |
type prover_result = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
57 |
{outcome: failure option, |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40202
diff
changeset
|
58 |
used_facts: (string * locality) list, |
40062 | 59 |
run_time_in_msecs: int option, |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
60 |
preplay: unit -> play, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
61 |
message: play -> string} |
39493
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents:
39492
diff
changeset
|
62 |
|
43051 | 63 |
type prover = |
64 |
params -> (string -> minimize_command) -> prover_problem -> prover_result |
|
35867 | 65 |
|
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
66 |
val dest_dir : string Config.T |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
67 |
val problem_prefix : string Config.T |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
68 |
val measure_run_time : bool Config.T |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
69 |
val atp_readable_names : bool Config.T |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
70 |
val smt_triggers : bool Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
71 |
val smt_weights : bool Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
72 |
val smt_weight_min_facts : int Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
73 |
val smt_min_weight : int Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
74 |
val smt_max_weight : int Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
75 |
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:
41242
diff
changeset
|
76 |
val smt_weight_curve : (int -> int) Unsynchronized.ref |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
77 |
val smt_max_slices : int Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
78 |
val smt_slice_fact_frac : real Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
79 |
val smt_slice_time_frac : real Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
80 |
val smt_slice_min_secs : int Config.T |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43004
diff
changeset
|
81 |
val das_tool : string |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
82 |
val select_smt_solver : string -> Proof.context -> Proof.context |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
83 |
val is_metis_prover : string -> bool |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
84 |
val is_atp : theory -> string -> bool |
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
85 |
val is_smt_prover : Proof.context -> string -> bool |
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
86 |
val is_unit_equational_atp : Proof.context -> string -> bool |
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
87 |
val is_prover_supported : Proof.context -> string -> bool |
40072
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents:
40071
diff
changeset
|
88 |
val is_prover_installed : Proof.context -> string -> bool |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
89 |
val default_max_relevant_for_prover : Proof.context -> bool -> string -> int |
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
90 |
val is_unit_equality : term -> bool |
42952
96f62b77748f
tuning -- the "appropriate" terminology is inspired from TPTP
blanchet
parents:
42944
diff
changeset
|
91 |
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:
40341
diff
changeset
|
92 |
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:
41335
diff
changeset
|
93 |
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:
41066
diff
changeset
|
94 |
val atp_relevance_fudge : relevance_fudge |
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
95 |
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:
40723
diff
changeset
|
96 |
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:
41242
diff
changeset
|
97 |
val weight_smt_fact : |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
98 |
Proof.context -> int -> ((string * locality) * thm) * int |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
99 |
-> (string * locality) * (int option * thm) |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
100 |
val untranslated_fact : prover_fact -> (string * locality) * thm |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
101 |
val smt_weighted_fact : |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
102 |
Proof.context -> int -> prover_fact * int |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
103 |
-> (string * locality) * (int option * thm) |
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
104 |
val supported_provers : Proof.context -> unit |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
105 |
val kill_provers : unit -> unit |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
106 |
val running_provers : unit -> unit |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
107 |
val messages : int option -> unit |
43033 | 108 |
val filter_used_facts : ''a list -> (''a * 'b) list -> (''a * 'b) list |
43021 | 109 |
val get_prover : Proof.context -> mode -> string -> prover |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
110 |
end; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
111 |
|
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
112 |
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
113 |
struct |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
114 |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
115 |
open ATP_Util |
38028 | 116 |
open ATP_Problem |
39491
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents:
39453
diff
changeset
|
117 |
open ATP_Proof |
38028 | 118 |
open ATP_Systems |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
119 |
open ATP_Translate |
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
120 |
open ATP_Reconstruct |
38023 | 121 |
open Sledgehammer_Util |
38988 | 122 |
open Sledgehammer_Filter |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
123 |
|
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
124 |
(** 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:
37581
diff
changeset
|
125 |
|
43021 | 126 |
datatype mode = Auto_Try | Try | Normal | Minimize |
127 |
||
38102
019a49759829
fix bug in the newly introduced "bound concealing" code
blanchet
parents:
38100
diff
changeset
|
128 |
(* Identifier to distinguish Sledgehammer from other tools using |
019a49759829
fix bug in the newly introduced "bound concealing" code
blanchet
parents:
38100
diff
changeset
|
129 |
"Async_Manager". *) |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43004
diff
changeset
|
130 |
val das_tool = "Sledgehammer" |
37585 | 131 |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
132 |
val metis_prover_names = [Metis_Tactics.metisN, Metis_Tactics.metisFT_N] |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
133 |
|
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
134 |
val is_metis_prover = member (op =) metis_prover_names |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
135 |
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:
43044
diff
changeset
|
136 |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
137 |
val select_smt_solver = |
41432
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41336
diff
changeset
|
138 |
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:
41241
diff
changeset
|
139 |
|
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
140 |
fun is_smt_prover ctxt name = |
41432
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41336
diff
changeset
|
141 |
member (op =) (SMT_Solver.available_solvers_of ctxt) name |
40062 | 142 |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
143 |
fun is_unit_equational_atp ctxt name = |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
144 |
let val thy = Proof_Context.theory_of ctxt in |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
145 |
case try (get_atp thy) name of |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
146 |
SOME {formats, ...} => member (op =) formats CNF_UEQ |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
147 |
| NONE => false |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
148 |
end |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
149 |
|
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
150 |
fun is_prover_supported ctxt name = |
42361 | 151 |
let val thy = Proof_Context.theory_of ctxt in |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
152 |
is_metis_prover name orelse is_atp thy name orelse is_smt_prover ctxt name |
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
153 |
end |
40072
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents:
40071
diff
changeset
|
154 |
|
41432
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41336
diff
changeset
|
155 |
fun is_prover_installed ctxt = |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
156 |
is_metis_prover orf is_smt_prover ctxt orf |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
157 |
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:
40723
diff
changeset
|
158 |
|
43063 | 159 |
fun get_slices num_facts slicing slices = |
160 |
(0 upto length slices - 1) ~~ slices |
|
161 |
|> (if slicing andalso |
|
162 |
exists (fn (_, (_, (max_facts, _))) => max_facts < num_facts) |
|
163 |
slices then |
|
164 |
I |
|
165 |
else |
|
166 |
List.last #> single) |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
167 |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
168 |
val metis_default_max_relevant = 20 |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
169 |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
170 |
fun default_max_relevant_for_prover ctxt slicing name = |
42361 | 171 |
let val thy = Proof_Context.theory_of ctxt in |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
172 |
if is_metis_prover name then |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
173 |
metis_default_max_relevant |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
174 |
else if is_atp thy name then |
42723 | 175 |
fold (Integer.max o fst o snd o snd o snd) |
43063 | 176 |
(get_slices 16384 (* large number *) slicing |
177 |
(#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:
43044
diff
changeset
|
178 |
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:
43044
diff
changeset
|
179 |
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:
40723
diff
changeset
|
180 |
end |
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
181 |
|
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
182 |
(* These are either simplified away by "Meson.presimplify" (most of the time) or |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
183 |
handled specially via "fFalse", "fTrue", ..., "fequal". *) |
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:
40070
diff
changeset
|
184 |
val atp_irrelevant_consts = |
41140
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
185 |
[@{const_name False}, @{const_name True}, @{const_name Not}, |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
186 |
@{const_name conj}, @{const_name disj}, @{const_name implies}, |
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
blanchet
parents:
41138
diff
changeset
|
187 |
@{const_name HOL.eq}, @{const_name If}, @{const_name Let}] |
40206 | 188 |
|
42956
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
189 |
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:
42952
diff
changeset
|
190 |
| 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:
42952
diff
changeset
|
191 |
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
192 |
(* 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:
42952
diff
changeset
|
193 |
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:
42952
diff
changeset
|
194 |
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:
42952
diff
changeset
|
195 |
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:
42952
diff
changeset
|
196 |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
197 |
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:
42943
diff
changeset
|
198 |
| is_unit_equality (Const (@{const_name all}, _) $ Abs (_, _, t)) = |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
199 |
is_unit_equality t |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
200 |
| is_unit_equality (Const (@{const_name All}, _) $ Abs (_, _, t)) = |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
201 |
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:
42952
diff
changeset
|
202 |
| 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:
42952
diff
changeset
|
203 |
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:
42952
diff
changeset
|
204 |
| 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:
42952
diff
changeset
|
205 |
is_good_unit_equality T t u |
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
206 |
| is_unit_equality _ = false |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
207 |
|
42952
96f62b77748f
tuning -- the "appropriate" terminology is inspired from TPTP
blanchet
parents:
42944
diff
changeset
|
208 |
fun is_appropriate_prop_for_prover ctxt name = |
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
209 |
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:
42943
diff
changeset
|
210 |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
211 |
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:
41241
diff
changeset
|
212 |
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:
41335
diff
changeset
|
213 |
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:
41335
diff
changeset
|
214 |
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:
41335
diff
changeset
|
215 |
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:
41335
diff
changeset
|
216 |
(true, []) |
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
217 |
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:
41335
diff
changeset
|
218 |
(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:
41335
diff
changeset
|
219 |
else |
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
220 |
(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:
41335
diff
changeset
|
221 |
end |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
222 |
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:
41335
diff
changeset
|
223 |
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:
40070
diff
changeset
|
224 |
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
225 |
(* FUDGE *) |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
226 |
val atp_relevance_fudge = |
42738 | 227 |
{local_const_multiplier = 1.5, |
41159
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
228 |
worse_irrel_freq = 100.0, |
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
229 |
higher_order_irrel_weight = 1.05, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
230 |
abs_rel_weight = 0.5, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
231 |
abs_irrel_weight = 2.0, |
42746
7e227e9de779
further lower penalty associated with existentials in Sledgehammer's relevance filter, so that "exhaust" facts are picked up
blanchet
parents:
42740
diff
changeset
|
232 |
skolem_irrel_weight = 0.25, |
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
233 |
theory_const_rel_weight = 0.5, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
234 |
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:
42730
diff
changeset
|
235 |
chained_const_irrel_weight = 0.25, |
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
236 |
intro_bonus = 0.15, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
237 |
elim_bonus = 0.15, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
238 |
simp_bonus = 0.15, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
239 |
local_bonus = 0.55, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
240 |
assum_bonus = 1.05, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
241 |
chained_bonus = 1.5, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
242 |
max_imperfect = 11.5, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
243 |
max_imperfect_exp = 1.0, |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
244 |
threshold_divisor = 2.0, |
41093 | 245 |
ridiculous_threshold = 0.01} |
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
246 |
|
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:
40070
diff
changeset
|
247 |
(* FUDGE (FIXME) *) |
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
248 |
val smt_relevance_fudge = |
42738 | 249 |
{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:
41152
diff
changeset
|
250 |
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:
40070
diff
changeset
|
251 |
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:
40070
diff
changeset
|
252 |
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:
40070
diff
changeset
|
253 |
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:
40070
diff
changeset
|
254 |
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:
40070
diff
changeset
|
255 |
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:
40070
diff
changeset
|
256 |
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:
42730
diff
changeset
|
257 |
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:
40070
diff
changeset
|
258 |
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:
40070
diff
changeset
|
259 |
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:
40070
diff
changeset
|
260 |
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:
40070
diff
changeset
|
261 |
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:
40070
diff
changeset
|
262 |
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:
40070
diff
changeset
|
263 |
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:
40070
diff
changeset
|
264 |
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:
40070
diff
changeset
|
265 |
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:
40070
diff
changeset
|
266 |
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:
40070
diff
changeset
|
267 |
ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge} |
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
268 |
|
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
269 |
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:
40723
diff
changeset
|
270 |
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:
40069
diff
changeset
|
271 |
|
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
272 |
fun supported_provers ctxt = |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
273 |
let |
42361 | 274 |
val thy = Proof_Context.theory_of ctxt |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
275 |
val (remote_provers, local_provers) = |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
276 |
metis_prover_names @ |
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
277 |
sort_strings (supported_atps thy) @ |
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
278 |
sort_strings (SMT_Solver.available_solvers_of ctxt) |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
279 |
|> List.partition (String.isPrefix remote_prefix) |
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
280 |
in |
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
281 |
Output.urgent_message ("Supported provers: " ^ |
40205
277508b07418
if "debug" is on, print list of relevant facts (poweruser request);
blanchet
parents:
40204
diff
changeset
|
282 |
commas (local_provers @ remote_provers) ^ ".") |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
283 |
end |
35969 | 284 |
|
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43004
diff
changeset
|
285 |
fun kill_provers () = Async_Manager.kill_threads das_tool "prover" |
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43004
diff
changeset
|
286 |
fun running_provers () = Async_Manager.running_threads das_tool "prover" |
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43004
diff
changeset
|
287 |
val messages = Async_Manager.thread_messages das_tool "prover" |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
288 |
|
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
289 |
(** problems, results, ATPs, etc. **) |
35969 | 290 |
|
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
291 |
datatype rich_type_system = |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
292 |
Dumb_Type_Sys of type_system | |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
293 |
Smart_Type_Sys of bool |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
294 |
|
35969 | 295 |
type params = |
41208
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
41171
diff
changeset
|
296 |
{debug: bool, |
35969 | 297 |
verbose: bool, |
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36064
diff
changeset
|
298 |
overlord: bool, |
41208
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
41171
diff
changeset
|
299 |
blocking: bool, |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
300 |
provers: string list, |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
301 |
type_sys: rich_type_system, |
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
42100
diff
changeset
|
302 |
relevance_thresholds: real * real, |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
42100
diff
changeset
|
303 |
max_relevant: int option, |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
304 |
max_mono_iters: int, |
42740
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents:
42738
diff
changeset
|
305 |
max_new_mono_instances: int, |
43063 | 306 |
explicit_apply: bool option, |
35969 | 307 |
isar_proof: bool, |
36924 | 308 |
isar_shrink_factor: int, |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
309 |
slicing: bool, |
38985
162bbbea4e4d
added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents:
38982
diff
changeset
|
310 |
timeout: Time.time, |
43015
21b6baec55b1
renamed "metis_timeout" to "preplay_timeout" and continued implementation
blanchet
parents:
43011
diff
changeset
|
311 |
preplay_timeout: Time.time, |
38985
162bbbea4e4d
added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents:
38982
diff
changeset
|
312 |
expect: string} |
35867 | 313 |
|
41090 | 314 |
datatype prover_fact = |
315 |
Untranslated_Fact of (string * locality) * thm | |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
316 |
SMT_Weighted_Fact of (string * locality) * (int option * thm) |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
317 |
|
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
318 |
type prover_problem = |
39318 | 319 |
{state: Proof.state, |
38998 | 320 |
goal: thm, |
321 |
subgoal: int, |
|
40065 | 322 |
subgoal_count: int, |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
323 |
facts: prover_fact list, |
41741 | 324 |
smt_filter: (string * locality) SMT_Solver.smt_filter_data option} |
35867 | 325 |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
326 |
type prover_result = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
327 |
{outcome: failure option, |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40202
diff
changeset
|
328 |
used_facts: (string * locality) list, |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
329 |
run_time_in_msecs: int option, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
330 |
preplay: unit -> play, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
331 |
message: play -> string} |
35867 | 332 |
|
43051 | 333 |
type prover = |
334 |
params -> (string -> minimize_command) -> prover_problem -> prover_result |
|
35867 | 335 |
|
43033 | 336 |
|
38023 | 337 |
(* configuration attributes *) |
338 |
||
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
339 |
(* 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:
42593
diff
changeset
|
340 |
val dest_dir = |
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42593
diff
changeset
|
341 |
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:
42593
diff
changeset
|
342 |
val problem_prefix = |
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42593
diff
changeset
|
343 |
Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob") |
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42593
diff
changeset
|
344 |
val measure_run_time = |
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42593
diff
changeset
|
345 |
Attrib.setup_config_bool @{binding sledgehammer_measure_run_time} (K false) |
28484 | 346 |
|
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
347 |
(* In addition to being easier to read, readable names are often much shorter, |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
348 |
especially if types are mangled in names. For these reason, they are enabled |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
349 |
by default. *) |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
350 |
val atp_readable_names = |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
351 |
Attrib.setup_config_bool @{binding sledgehammer_atp_readable_names} (K true) |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
352 |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
353 |
val smt_triggers = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
354 |
Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
355 |
val smt_weights = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
356 |
Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
357 |
val smt_weight_min_facts = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
358 |
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:
41242
diff
changeset
|
359 |
|
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
360 |
(* FUDGE *) |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
361 |
val smt_min_weight = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
362 |
Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
363 |
val smt_max_weight = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
364 |
Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
365 |
val smt_max_weight_index = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
366 |
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:
41242
diff
changeset
|
367 |
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:
41242
diff
changeset
|
368 |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
369 |
fun smt_fact_weight ctxt j num_facts = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
370 |
if Config.get ctxt smt_weights andalso |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
371 |
num_facts >= Config.get ctxt smt_weight_min_facts then |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
372 |
let |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
373 |
val min = Config.get ctxt smt_min_weight |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
374 |
val max = Config.get ctxt smt_max_weight |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
375 |
val max_index = Config.get ctxt smt_max_weight_index |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
376 |
val curve = !smt_weight_curve |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
377 |
in |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
378 |
SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1)) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
379 |
div curve max_index) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
380 |
end |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
381 |
else |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
382 |
NONE |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
383 |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
384 |
fun weight_smt_fact ctxt num_facts ((info, th), j) = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
385 |
let val thy = Proof_Context.theory_of ctxt in |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
386 |
(info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy)) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
387 |
end |
38023 | 388 |
|
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
389 |
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:
41241
diff
changeset
|
390 |
| untranslated_fact (SMT_Weighted_Fact (info, (_, th))) = (info, th) |
42994
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
blanchet
parents:
42972
diff
changeset
|
391 |
fun atp_translated_fact ctxt format type_sys fact = |
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
blanchet
parents:
42972
diff
changeset
|
392 |
translate_atp_fact ctxt format type_sys false (untranslated_fact fact) |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
393 |
fun smt_weighted_fact _ _ (SMT_Weighted_Fact p, _) = p |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
394 |
| smt_weighted_fact ctxt num_facts (fact, j) = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
395 |
(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:
41242
diff
changeset
|
396 |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
397 |
fun overlord_file_location_for_prover prover = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
398 |
(getenv "ISABELLE_HOME_USER", "prob_" ^ prover) |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
399 |
|
43033 | 400 |
fun with_path cleanup after f path = |
401 |
Exn.capture f path |
|
402 |
|> tap (fn _ => cleanup path) |
|
403 |
|> Exn.release |
|
404 |
|> tap (after path) |
|
405 |
||
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
406 |
fun proof_banner mode name = |
43033 | 407 |
case mode of |
408 |
Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof" |
|
409 |
| Try => "Sledgehammer (" ^ quote name ^ ") found a proof" |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
410 |
| _ => "Try this" |
43033 | 411 |
|
412 |
(* based on "Mirabelle.can_apply" and generalized *) |
|
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
413 |
fun timed_apply timeout tac state i = |
43033 | 414 |
let |
415 |
val {context = ctxt, facts, goal} = Proof.goal state |
|
416 |
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:
43033
diff
changeset
|
417 |
in TimeLimit.timeLimit timeout (try (Seq.pull o full_tac)) goal end |
43033 | 418 |
|
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
419 |
fun tac_for_reconstructor Metis = Metis_Tactics.metisHO_tac |
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
420 |
| tac_for_reconstructor MetisFT = Metis_Tactics.metisFT_tac |
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
421 |
| tac_for_reconstructor _ = raise Fail "unexpected reconstructor" |
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
422 |
|
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
423 |
fun timed_reconstructor reconstructor debug timeout ths = |
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
424 |
(Config.put Metis_Tactics.verbose debug |
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
425 |
#> (fn ctxt => tac_for_reconstructor reconstructor ctxt ths)) |
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
426 |
|> timed_apply timeout |
43033 | 427 |
|
428 |
fun filter_used_facts used = filter (member (op =) used o fst) |
|
429 |
||
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
430 |
fun play_one_line_proof debug timeout ths state i reconstructors = |
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
431 |
let |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
432 |
fun play [] [] = Failed_to_Play |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
433 |
| play (timed_out :: _) [] = Trust_Playable (timed_out, SOME timeout) |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
434 |
| play timed_out (reconstructor :: reconstructors) = |
43044
5945375700aa
always check plain "metis" even if the ATP proof seems to require "metisFT" -- maybe the proof is needlessly complicated
blanchet
parents:
43037
diff
changeset
|
435 |
let val timer = Timer.startRealTimer () in |
5945375700aa
always check plain "metis" even if the ATP proof seems to require "metisFT" -- maybe the proof is needlessly complicated
blanchet
parents:
43037
diff
changeset
|
436 |
case timed_reconstructor reconstructor debug timeout ths state i of |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
437 |
SOME (SOME _) => Played (reconstructor, Timer.checkRealTimer timer) |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
438 |
| _ => play timed_out reconstructors |
43044
5945375700aa
always check plain "metis" even if the ATP proof seems to require "metisFT" -- maybe the proof is needlessly complicated
blanchet
parents:
43037
diff
changeset
|
439 |
end |
5945375700aa
always check plain "metis" even if the ATP proof seems to require "metisFT" -- maybe the proof is needlessly complicated
blanchet
parents:
43037
diff
changeset
|
440 |
handle TimeLimit.TimeOut => |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
441 |
play (reconstructor :: timed_out) reconstructors |
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
442 |
in |
43044
5945375700aa
always check plain "metis" even if the ATP proof seems to require "metisFT" -- maybe the proof is needlessly complicated
blanchet
parents:
43037
diff
changeset
|
443 |
if timeout = Time.zeroTime then Trust_Playable (hd reconstructors, NONE) |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
444 |
else play [] reconstructors |
43033 | 445 |
end |
446 |
||
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
447 |
|
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
448 |
(* generic TPTP-based ATPs *) |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
449 |
|
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:
42729
diff
changeset
|
450 |
(* 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:
42729
diff
changeset
|
451 |
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:
42729
diff
changeset
|
452 |
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:
42729
diff
changeset
|
453 |
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:
42729
diff
changeset
|
454 |
|
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:
42729
diff
changeset
|
455 |
(* 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:
42729
diff
changeset
|
456 |
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:
42729
diff
changeset
|
457 |
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:
42729
diff
changeset
|
458 |
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:
42729
diff
changeset
|
459 |
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:
42729
diff
changeset
|
460 |
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:
42729
diff
changeset
|
461 |
| 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:
42729
diff
changeset
|
462 |
| 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:
42729
diff
changeset
|
463 |
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:
42729
diff
changeset
|
464 |
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:
42729
diff
changeset
|
465 |
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:
42729
diff
changeset
|
466 |
(_, @{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:
42729
diff
changeset
|
467 |
| (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:
42729
diff
changeset
|
468 |
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:
42729
diff
changeset
|
469 |
| (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:
42729
diff
changeset
|
470 |
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:
42729
diff
changeset
|
471 |
| (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:
42729
diff
changeset
|
472 |
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:
42729
diff
changeset
|
473 |
| (_, @{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:
42729
diff
changeset
|
474 |
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:
42729
diff
changeset
|
475 |
(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:
42729
diff
changeset
|
476 |
| (_, @{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:
42729
diff
changeset
|
477 |
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:
42729
diff
changeset
|
478 |
(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:
42729
diff
changeset
|
479 |
| (_, @{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:
42729
diff
changeset
|
480 |
| (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:
42729
diff
changeset
|
481 |
| (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:
42729
diff
changeset
|
482 |
| (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:
42729
diff
changeset
|
483 |
| (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:
42729
diff
changeset
|
484 |
| _ => 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:
42729
diff
changeset
|
485 |
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:
42729
diff
changeset
|
486 |
|
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:
42729
diff
changeset
|
487 |
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:
42729
diff
changeset
|
488 |
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:
42729
diff
changeset
|
489 |
| 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:
42729
diff
changeset
|
490 |
| _ => 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:
42729
diff
changeset
|
491 |
|
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:
42729
diff
changeset
|
492 |
(* 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:
42729
diff
changeset
|
493 |
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:
42729
diff
changeset
|
494 |
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:
42729
diff
changeset
|
495 |
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:
42729
diff
changeset
|
496 |
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:
42729
diff
changeset
|
497 |
|
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:
42729
diff
changeset
|
498 |
(* 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:
42729
diff
changeset
|
499 |
"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:
42729
diff
changeset
|
500 |
are omitted. *) |
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
501 |
fun is_dangerous_prop ctxt = |
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
502 |
transform_elim_prop |
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:
42729
diff
changeset
|
503 |
#> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf |
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:
42729
diff
changeset
|
504 |
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:
42729
diff
changeset
|
505 |
|
40409
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents:
40370
diff
changeset
|
506 |
fun int_opt_add (SOME m) (SOME n) = SOME (m + n) |
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents:
40370
diff
changeset
|
507 |
| int_opt_add _ _ = NONE |
40062 | 508 |
|
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
509 |
val atp_blacklist_max_iters = 10 |
39492 | 510 |
(* Important messages are important but not so important that users want to see |
511 |
them each time. *) |
|
42609 | 512 |
val atp_important_message_keep_quotient = 10 |
39492 | 513 |
|
42589
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42588
diff
changeset
|
514 |
val fallback_best_type_systems = |
42853
de1fe9eaf3f4
tweaked ATP type systems further based on Judgment Day
blanchet
parents:
42849
diff
changeset
|
515 |
[Preds (Mangled_Monomorphic, Nonmonotonic_Types, Light)] |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
516 |
|
42722 | 517 |
fun adjust_dumb_type_sys formats (Simple_Types level) = |
42963 | 518 |
if member (op =) formats THF then |
519 |
(THF, Simple_Types level) |
|
520 |
else if member (op =) formats TFF then |
|
521 |
(TFF, Simple_Types level) |
|
522 |
else |
|
523 |
adjust_dumb_type_sys formats (Preds (Mangled_Monomorphic, level, Heavy)) |
|
42849
ba45312308b5
reintroduce TFF workaround for limitations of actual TFF implementations (ToFoF, SNARK)
blanchet
parents:
42837
diff
changeset
|
524 |
| adjust_dumb_type_sys formats type_sys = |
42998
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
525 |
(case hd formats of |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
526 |
CNF_UEQ => |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
527 |
(CNF_UEQ, case type_sys of |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
528 |
Preds stuff => |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
529 |
(if is_type_sys_fairly_sound type_sys then Preds else Tags) |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
530 |
stuff |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
531 |
| _ => type_sys) |
1c80902d0456
fully support all type system encodings in typed formats (TFF, THF)
blanchet
parents:
42994
diff
changeset
|
532 |
| format => (format, type_sys)) |
42994
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
blanchet
parents:
42972
diff
changeset
|
533 |
|
43051 | 534 |
fun choose_format_and_type_sys _ formats (Dumb_Type_Sys type_sys) = |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42578
diff
changeset
|
535 |
adjust_dumb_type_sys formats type_sys |
43051 | 536 |
| choose_format_and_type_sys best_type_systems formats |
537 |
(Smart_Type_Sys full_types) = |
|
42613
23b13b1bd565
use strings to encode type systems in ATP module, to reduce the amount of out-of-place information and also to make it easier to print the type system used
blanchet
parents:
42609
diff
changeset
|
538 |
map type_sys_from_string best_type_systems @ fallback_best_type_systems |
42589
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42588
diff
changeset
|
539 |
|> find_first (if full_types then is_type_sys_virtually_sound else K true) |
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42588
diff
changeset
|
540 |
|> the |> adjust_dumb_type_sys formats |
42548
ea2a28b1938f
make sure the minimizer monomorphizes when it should
blanchet
parents:
42544
diff
changeset
|
541 |
|
43051 | 542 |
val metis_minimize_max_time = seconds 2.0 |
543 |
||
544 |
fun choose_minimize_command minimize_command name preplay = |
|
545 |
(case preplay of |
|
546 |
Played (reconstructor, time) => |
|
547 |
if Time.<= (time, metis_minimize_max_time) then |
|
548 |
reconstructor_name reconstructor |
|
549 |
else |
|
550 |
name |
|
551 |
| _ => name) |
|
552 |
|> minimize_command |
|
553 |
||
42740
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents:
42738
diff
changeset
|
554 |
fun repair_smt_monomorph_context debug max_mono_iters max_mono_instances = |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
555 |
Config.put SMT_Config.verbose debug |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
556 |
#> Config.put SMT_Config.monomorph_full false |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
557 |
#> Config.put SMT_Config.monomorph_limit max_mono_iters |
42740
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents:
42738
diff
changeset
|
558 |
#> Config.put SMT_Config.monomorph_instances max_mono_instances |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
559 |
|
43021 | 560 |
fun run_atp mode name |
42578
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
blanchet
parents:
42577
diff
changeset
|
561 |
({exec, required_execs, arguments, proof_delims, known_failures, |
42723 | 562 |
conj_sym_kind, prem_kind, formats, best_slices, ...} : atp_config) |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
563 |
({debug, verbose, overlord, type_sys, max_relevant, max_mono_iters, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
564 |
max_new_mono_instances, explicit_apply, isar_proof, |
43033 | 565 |
isar_shrink_factor, slicing, timeout, preplay_timeout, ...} : params) |
43037 | 566 |
minimize_command |
567 |
({state, goal, subgoal, subgoal_count, facts, ...} : prover_problem) = |
|
38023 | 568 |
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:
42181
diff
changeset
|
569 |
val thy = Proof.theory_of state |
39318 | 570 |
val ctxt = Proof.context_of state |
43004
20e9caff1f86
fix soundness bug in Sledgehammer: distinguish params in goals from fixed variables in context
blanchet
parents:
42998
diff
changeset
|
571 |
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:
41152
diff
changeset
|
572 |
val (dest_dir, problem_prefix) = |
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
573 |
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:
41152
diff
changeset
|
574 |
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:
40060
diff
changeset
|
575 |
val problem_file_name = |
41159
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
576 |
Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^ |
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
577 |
"_" ^ string_of_int subgoal) |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
578 |
val problem_path_name = |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
579 |
if dest_dir = "" then |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
580 |
File.tmp_path problem_file_name |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
581 |
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:
40060
diff
changeset
|
582 |
Path.append (Path.explode dest_dir) problem_file_name |
39003
c2aebd79981f
run relevance filter in a thread, to avoid blocking
blanchet
parents:
39000
diff
changeset
|
583 |
else |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
584 |
error ("No such directory: " ^ quote dest_dir ^ ".") |
39003
c2aebd79981f
run relevance filter in a thread, to avoid blocking
blanchet
parents:
39000
diff
changeset
|
585 |
val measure_run_time = verbose orelse Config.get ctxt measure_run_time |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38091
diff
changeset
|
586 |
val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec) |
38023 | 587 |
fun split_time s = |
588 |
let |
|
42448 | 589 |
val split = String.tokens (fn c => str c = "\n") |
590 |
val (output, t) = s |> split |> split_last |> apfst cat_lines |
|
591 |
fun as_num f = f >> (fst o read_int) |
|
592 |
val num = as_num (Scan.many1 Symbol.is_ascii_digit) |
|
593 |
val digit = Scan.one Symbol.is_ascii_digit |
|
594 |
val num3 = as_num (digit ::: digit ::: (digit >> single)) |
|
595 |
val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b) |
|
40627
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents:
40562
diff
changeset
|
596 |
val as_time = Scan.read Symbol.stopper time o raw_explode |
42448 | 597 |
in (output, as_time t) end |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
598 |
fun run_on prob_file = |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38091
diff
changeset
|
599 |
case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of |
38032 | 600 |
(home_var, _) :: _ => |
38023 | 601 |
error ("The environment variable " ^ quote home_var ^ " is not set.") |
38032 | 602 |
| [] => |
603 |
if File.exists command then |
|
604 |
let |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
605 |
(* If slicing is disabled, we expand the last slice to fill the |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
606 |
entire time available. *) |
43063 | 607 |
val actual_slices = |
608 |
get_slices (length facts) slicing (best_slices ctxt) |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
609 |
val num_actual_slices = length actual_slices |
42445
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
610 |
fun monomorphize_facts facts = |
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
611 |
let |
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
612 |
val facts = facts |> map untranslated_fact |
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
613 |
(* pseudo-theorem involving the same constants as the subgoal *) |
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
614 |
val subgoal_th = |
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
615 |
Logic.list_implies (hyp_ts, concl_t) |
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
616 |
|> Skip_Proof.make_thm thy |
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
617 |
val indexed_facts = |
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
618 |
(~1, subgoal_th) :: (0 upto length facts - 1 ~~ map snd facts) |
42740
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents:
42738
diff
changeset
|
619 |
val max_mono_instances = max_new_mono_instances + length facts |
42445
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
620 |
in |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
621 |
ctxt |> repair_smt_monomorph_context debug max_mono_iters |
42740
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents:
42738
diff
changeset
|
622 |
max_mono_instances |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
623 |
|> SMT_Monomorph.monomorph indexed_facts |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
624 |
|> fst |> sort (int_ord o pairself fst) |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
625 |
|> filter_out (curry (op =) ~1 o fst) |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
626 |
|> map (Untranslated_Fact o apfst (fst o nth facts)) |
42445
c6ea64ebb8c5
fixed interaction between monomorphization and slicing for ATPs
blanchet
parents:
42444
diff
changeset
|
627 |
end |
42723 | 628 |
fun run_slice blacklist (slice, (time_frac, (complete, |
629 |
(best_max_relevant, best_type_systems)))) |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
630 |
time_left = |
38032 | 631 |
let |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
632 |
val num_facts = |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
633 |
length facts |> is_none max_relevant |
42723 | 634 |
? Integer.min best_max_relevant |
635 |
val (format, type_sys) = |
|
43051 | 636 |
choose_format_and_type_sys best_type_systems formats type_sys |
42638
a7a30721767a
have each ATP filter out dangerous facts for themselves, based on their type system
blanchet
parents:
42623
diff
changeset
|
637 |
val fairly_sound = is_type_sys_fairly_sound type_sys |
42451
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents:
42450
diff
changeset
|
638 |
val facts = |
42671
390de893659a
fixed per-ATP dangerous axiom detection -- embarrassing bugs introduced in change a7a30721767a
blanchet
parents:
42647
diff
changeset
|
639 |
facts |> not fairly_sound |
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
640 |
? filter_out (is_dangerous_prop ctxt o prop_of o snd |
42638
a7a30721767a
have each ATP filter out dangerous facts for themselves, based on their type system
blanchet
parents:
42623
diff
changeset
|
641 |
o untranslated_fact) |
a7a30721767a
have each ATP filter out dangerous facts for themselves, based on their type system
blanchet
parents:
42623
diff
changeset
|
642 |
|> take num_facts |
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
643 |
|> not (null blacklist) |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
644 |
? filter_out (member (op =) blacklist o fst |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
645 |
o untranslated_fact) |
42589
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42588
diff
changeset
|
646 |
|> polymorphism_of_type_sys type_sys <> Polymorphic |
9f7c48463645
restructured type systems some more -- the old naming schemes had "argshg diff |less" and "tagshg diff |less" as equivalent and didn't support a monomorphic version of "tags"
blanchet
parents:
42588
diff
changeset
|
647 |
? monomorphize_facts |
42994
fe291ab75eb5
towards supporting non-simply-typed encodings for TFF and THF (for orthogonality and experiments)
blanchet
parents:
42972
diff
changeset
|
648 |
|> map (atp_translated_fact ctxt format type_sys) |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
649 |
val real_ms = Real.fromInt o Time.toMilliseconds |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
650 |
val slice_timeout = |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
651 |
((real_ms time_left |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
652 |
|> (if slice < num_actual_slices - 1 then |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
653 |
curry Real.min (time_frac * real_ms timeout) |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
654 |
else |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
655 |
I)) |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
656 |
* 0.001) |> seconds |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
657 |
val _ = |
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
658 |
if debug then |
42699 | 659 |
quote name ^ " slice #" ^ string_of_int (slice + 1) ^ |
660 |
" with " ^ string_of_int num_facts ^ " fact" ^ |
|
661 |
plural_s num_facts ^ " for " ^ |
|
662 |
string_from_time slice_timeout ^ "..." |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
663 |
|> Output.urgent_message |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
664 |
else |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
665 |
() |
42541
8938507b2054
move type declarations to the front, for TFF-compliance
blanchet
parents:
42523
diff
changeset
|
666 |
val (atp_problem, pool, conjecture_offset, facts_offset, |
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
667 |
fact_names, typed_helpers, sym_tab) = |
42939 | 668 |
prepare_atp_problem ctxt format conj_sym_kind prem_kind |
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
669 |
type_sys explicit_apply |
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
670 |
(Config.get ctxt atp_readable_names) hyp_ts concl_t facts |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
671 |
fun weights () = atp_problem_weights atp_problem |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
672 |
val core = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
673 |
File.shell_path command ^ " " ^ |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
674 |
arguments ctxt slice slice_timeout weights ^ " " ^ |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
675 |
File.shell_path prob_file |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
676 |
val command = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
677 |
(if measure_run_time then |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
678 |
"TIMEFORMAT='%3R'; { time " ^ core ^ " ; }" |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
679 |
else |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
680 |
"exec " ^ core) ^ " 2>&1" |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
681 |
val _ = |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
682 |
atp_problem |
42709
e7af132d48fe
allow each prover to specify its own formula kind for symbols occurring in the conjecture
blanchet
parents:
42699
diff
changeset
|
683 |
|> tptp_strings_for_atp_problem format |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
684 |
|> cons ("% " ^ command ^ "\n") |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
685 |
|> File.write_list prob_file |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
686 |
val conjecture_shape = |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
687 |
conjecture_offset + 1 |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
688 |
upto conjecture_offset + length hyp_ts + 1 |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
689 |
|> map single |
38032 | 690 |
val ((output, msecs), res_code) = |
691 |
bash_output command |
|
692 |
|>> (if overlord then |
|
693 |
prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n") |
|
694 |
else |
|
695 |
I) |
|
40062 | 696 |
|>> (if measure_run_time then split_time else rpair NONE) |
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42448
diff
changeset
|
697 |
val (atp_proof, outcome) = |
42849
ba45312308b5
reintroduce TFF workaround for limitations of actual TFF implementations (ToFoF, SNARK)
blanchet
parents:
42837
diff
changeset
|
698 |
extract_tstplike_proof_and_outcome verbose complete res_code |
ba45312308b5
reintroduce TFF workaround for limitations of actual TFF implementations (ToFoF, SNARK)
blanchet
parents:
42837
diff
changeset
|
699 |
proof_delims known_failures output |
42943 | 700 |
|>> atp_proof_from_tstplike_proof atp_problem |
42965
1403595ec38c
slightly gracefuller handling of LEO-II and Satallax output
blanchet
parents:
42963
diff
changeset
|
701 |
handle UNRECOGNIZED_ATP_PROOF () => ([], SOME ProofIncomplete) |
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
702 |
val (conjecture_shape, facts_offset, fact_names, |
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
703 |
typed_helpers) = |
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42448
diff
changeset
|
704 |
if is_none outcome then |
42647
59142dbfa3ba
no need to generate fact numbers for polymorphic type systems -- this confuses the TPTP exporter
blanchet
parents:
42646
diff
changeset
|
705 |
repair_conjecture_shape_and_fact_names type_sys output |
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
706 |
conjecture_shape facts_offset fact_names typed_helpers |
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42448
diff
changeset
|
707 |
else |
42587
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents:
42579
diff
changeset
|
708 |
(* don't bother repairing *) |
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
709 |
(conjecture_shape, facts_offset, fact_names, typed_helpers) |
42449
494e4ac5b0f8
detect some unsound proofs before showing them to the user
blanchet
parents:
42448
diff
changeset
|
710 |
val outcome = |
42451
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents:
42450
diff
changeset
|
711 |
case outcome of |
42587
4fbb1de05169
fixed SPASS fact offset calculation and report unexpected unsound proofs with so-called sound encodings
blanchet
parents:
42579
diff
changeset
|
712 |
NONE => |
42968
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents:
42965
diff
changeset
|
713 |
used_facts_in_unsound_atp_proof ctxt type_sys |
42876
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
blanchet
parents:
42853
diff
changeset
|
714 |
conjecture_shape facts_offset fact_names atp_proof |
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
blanchet
parents:
42853
diff
changeset
|
715 |
|> Option.map (fn facts => |
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
blanchet
parents:
42853
diff
changeset
|
716 |
UnsoundProof (is_type_sys_virtually_sound type_sys, |
42943 | 717 |
facts |> sort string_ord)) |
42451
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents:
42450
diff
changeset
|
718 |
| SOME Unprovable => |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
719 |
if null blacklist then outcome else SOME GaveUp |
42451
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents:
42450
diff
changeset
|
720 |
| _ => outcome |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
721 |
in |
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
722 |
((pool, conjecture_shape, facts_offset, fact_names, |
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
723 |
typed_helpers, sym_tab), |
42723 | 724 |
(output, msecs, type_sys, atp_proof, outcome)) |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
725 |
end |
38645 | 726 |
val timer = Timer.startRealTimer () |
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
727 |
fun maybe_run_slice blacklist slice |
42723 | 728 |
(result as (_, (_, msecs0, _, _, SOME _))) = |
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
729 |
let |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
730 |
val time_left = Time.- (timeout, Timer.checkRealTimer timer) |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
731 |
in |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
732 |
if Time.<= (time_left, Time.zeroTime) then |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
733 |
result |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
734 |
else |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
735 |
(run_slice blacklist slice time_left |
42723 | 736 |
|> (fn (stuff, (output, msecs, type_sys, atp_proof, |
737 |
outcome)) => |
|
738 |
(stuff, (output, int_opt_add msecs0 msecs, type_sys, |
|
739 |
atp_proof, outcome)))) |
|
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
740 |
end |
42450
2765d4fb2b9c
automatically retry with full-types upon unsound proof
blanchet
parents:
42449
diff
changeset
|
741 |
| maybe_run_slice _ _ result = result |
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
742 |
fun maybe_blacklist_facts_and_retry iter blacklist |
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
743 |
(result as ((_, _, facts_offset, fact_names, _, _), |
42723 | 744 |
(_, _, type_sys, atp_proof, |
42876
e336ef6313aa
more informative message when Sledgehammer finds an unsound proof
blanchet
parents:
42853
diff
changeset
|
745 |
SOME (UnsoundProof (false, _))))) = |
42968
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents:
42965
diff
changeset
|
746 |
(case used_facts_in_atp_proof ctxt type_sys facts_offset |
74415622d293
more work on parsing LEO-II proofs and extracting uses of extensionality
blanchet
parents:
42965
diff
changeset
|
747 |
fact_names atp_proof of |
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
748 |
[] => result |
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
749 |
| new_baddies => |
42835 | 750 |
if slicing andalso iter < atp_blacklist_max_iters - 1 then |
42777 | 751 |
let val blacklist = new_baddies @ blacklist in |
752 |
result |
|
753 |
|> maybe_run_slice blacklist (List.last actual_slices) |
|
754 |
|> maybe_blacklist_facts_and_retry (iter + 1) blacklist |
|
755 |
end |
|
756 |
else |
|
757 |
result) |
|
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
758 |
| maybe_blacklist_facts_and_retry _ _ result = result |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
759 |
in |
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
760 |
((Symtab.empty, [], 0, Vector.fromList [], [], Symtab.empty), |
42723 | 761 |
("", SOME 0, hd fallback_best_type_systems, [], |
762 |
SOME InternalError)) |
|
42451
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents:
42450
diff
changeset
|
763 |
|> fold (maybe_run_slice []) actual_slices |
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents:
42450
diff
changeset
|
764 |
(* The ATP found an unsound proof? Automatically try again |
a75fcd103cbb
automatically remove offending facts when faced with an unsound proof -- instead of using the highly inefficient "full_types" option
blanchet
parents:
42450
diff
changeset
|
765 |
without the offending facts! *) |
42452
f7f796ce5d68
iterate the unsound-fact-set removal process to recover even more unsound proofs
blanchet
parents:
42451
diff
changeset
|
766 |
|> maybe_blacklist_facts_and_retry 0 [] |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
767 |
end |
38032 | 768 |
else |
41944
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
wenzelm
parents:
41799
diff
changeset
|
769 |
error ("Bad executable: " ^ Path.print command ^ ".") |
38023 | 770 |
|
771 |
(* If the problem file has not been exported, remove it; otherwise, export |
|
772 |
the proof file too. *) |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
773 |
fun cleanup prob_file = |
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
774 |
if dest_dir = "" then try File.rm prob_file else NONE |
42723 | 775 |
fun export prob_file (_, (output, _, _, _, _)) = |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
776 |
if dest_dir = "" then |
38023 | 777 |
() |
778 |
else |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
779 |
File.write (Path.explode (Path.implode prob_file ^ "_proof")) output |
42881
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
780 |
val ((pool, conjecture_shape, facts_offset, fact_names, typed_helpers, |
dbdfe2d5b6ab
automatically use "metisFT" when typed helpers are necessary
blanchet
parents:
42876
diff
changeset
|
781 |
sym_tab), |
42723 | 782 |
(output, msecs, type_sys, atp_proof, outcome)) = |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
783 |
with_path cleanup export run_on problem_path_name |
39492 | 784 |
val important_message = |
43021 | 785 |
if mode = Normal andalso |
42609 | 786 |
random_range 0 (atp_important_message_keep_quotient - 1) = 0 then |
39492 | 787 |
extract_important_message output |
788 |
else |
|
789 |
"" |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
790 |
val (used_facts, preplay, message) = |
38023 | 791 |
case outcome of |
792 |
NONE => |
|
43033 | 793 |
let |
794 |
val used_facts = |
|
795 |
used_facts_in_atp_proof ctxt type_sys facts_offset fact_names |
|
796 |
atp_proof |
|
797 |
in |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
798 |
(used_facts, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
799 |
fn () => |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
800 |
let |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
801 |
val used_ths = |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
802 |
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:
43051
diff
changeset
|
803 |
|> map snd |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
804 |
in |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
805 |
play_one_line_proof debug preplay_timeout used_ths state subgoal |
43063 | 806 |
[Metis, MetisFT] |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
807 |
end, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
808 |
fn preplay => |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
809 |
let |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
810 |
val full_types = uses_typed_helpers typed_helpers atp_proof |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
811 |
val isar_params = |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
812 |
(debug, full_types, isar_shrink_factor, type_sys, pool, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
813 |
conjecture_shape, facts_offset, fact_names, sym_tab, atp_proof, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
814 |
goal) |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
815 |
val one_line_params = |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
816 |
(preplay, proof_banner mode name, used_facts, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
817 |
choose_minimize_command minimize_command name preplay, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
818 |
subgoal, subgoal_count) |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
819 |
in |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
820 |
proof_text ctxt isar_proof isar_params one_line_params ^ |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
821 |
(if verbose then |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
822 |
"\nATP real CPU time: " ^ |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
823 |
string_from_time (Time.fromMilliseconds (the msecs)) ^ "." |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
824 |
else |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
825 |
"") ^ |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
826 |
(if important_message <> "" then |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
827 |
"\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
828 |
important_message |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
829 |
else |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
830 |
"") |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
831 |
end) |
43033 | 832 |
end |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
833 |
| SOME failure => |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
834 |
([], K Failed_to_Play, fn _ => string_for_failure failure) |
38023 | 835 |
in |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
836 |
{outcome = outcome, used_facts = used_facts, run_time_in_msecs = msecs, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
837 |
preplay = preplay, message = message} |
38023 | 838 |
end |
839 |
||
40669 | 840 |
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until |
841 |
these are sorted out properly in the SMT module, we have to interpret these |
|
842 |
ourselves. *) |
|
40684
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
843 |
val remote_smt_failures = |
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
844 |
[(22, CantConnect), |
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
845 |
(2, NoLibwwwPerl)] |
41222 | 846 |
val z3_wrapper_failures = |
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
847 |
[(11, InternalError), |
41222 | 848 |
(12, InternalError), |
849 |
(13, InternalError)] |
|
40684
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
850 |
val z3_failures = |
41236 | 851 |
[(101, OutOfResources), |
852 |
(103, MalformedInput), |
|
41222 | 853 |
(110, MalformedInput)] |
40684
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
854 |
val unix_failures = |
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
855 |
[(139, Crashed)] |
41222 | 856 |
val smt_failures = |
41799
98b3d5ce0935
exit code 127 can mean many things -- not just (and probably not) Perl missing
blanchet
parents:
41790
diff
changeset
|
857 |
remote_smt_failures @ z3_wrapper_failures @ z3_failures @ unix_failures |
40555 | 858 |
|
42100
062381c5f9f8
more precise failure reporting in Sledgehammer/SMT
blanchet
parents:
42061
diff
changeset
|
859 |
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:
43044
diff
changeset
|
860 |
if is_real_cex then Unprovable else GaveUp |
41222 | 861 |
| failure_from_smt_failure SMT_Failure.Time_Out = TimedOut |
862 |
| failure_from_smt_failure (SMT_Failure.Abnormal_Termination code) = |
|
863 |
(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:
40669
diff
changeset
|
864 |
SOME failure => failure |
41259
13972ced98d9
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
blanchet
parents:
41256
diff
changeset
|
865 |
| 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:
41256
diff
changeset
|
866 |
string_of_int code ^ ".")) |
41222 | 867 |
| failure_from_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources |
868 |
| 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:
42060
diff
changeset
|
869 |
UnknownError msg |
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
870 |
|
40698 | 871 |
(* FUDGE *) |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
872 |
val smt_max_slices = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
873 |
Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
874 |
val smt_slice_fact_frac = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
875 |
Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac} (K 0.5) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
876 |
val smt_slice_time_frac = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
877 |
Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.5) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
878 |
val smt_slice_min_secs = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
879 |
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:
40370
diff
changeset
|
880 |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
881 |
fun smt_filter_loop ctxt name |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
882 |
({debug, verbose, overlord, max_mono_iters, |
42740
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents:
42738
diff
changeset
|
883 |
max_new_mono_instances, timeout, slicing, ...} : params) |
41741 | 884 |
state i smt_filter = |
40409
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents:
40370
diff
changeset
|
885 |
let |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
886 |
val max_slices = if slicing 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:
41241
diff
changeset
|
887 |
val repair_context = |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
888 |
select_smt_solver name |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
889 |
#> (if overlord then |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
890 |
Config.put SMT_Config.debug_files |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
891 |
(overlord_file_location_for_prover name |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
892 |
|> (fn (path, name) => path ^ "/" ^ name)) |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
893 |
else |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
894 |
I) |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
895 |
#> 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:
41241
diff
changeset
|
896 |
val state = state |> Proof.map_context repair_context |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
897 |
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:
40471
diff
changeset
|
898 |
let |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
899 |
val timer = Timer.startRealTimer () |
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
900 |
val state = |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
901 |
state |> Proof.map_context |
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
902 |
(repair_smt_monomorph_context debug max_mono_iters |
42740
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents:
42738
diff
changeset
|
903 |
(max_new_mono_instances + length facts)) |
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
904 |
val ms = timeout |> Time.toMilliseconds |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
905 |
val slice_timeout = |
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
906 |
if slice < max_slices then |
41169 | 907 |
Int.min (ms, |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
908 |
Int.max (1000 * Config.get ctxt smt_slice_min_secs, |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
909 |
Real.ceil (Config.get ctxt smt_slice_time_frac |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
910 |
* Real.fromInt ms))) |
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
911 |
|> Time.fromMilliseconds |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
912 |
else |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
913 |
timeout |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
914 |
val num_facts = length facts |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
915 |
val _ = |
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
916 |
if debug then |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
917 |
quote name ^ " slice " ^ string_of_int slice ^ " with " ^ |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
918 |
string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ " for " ^ |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
919 |
string_from_time slice_timeout ^ "..." |
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
920 |
|> Output.urgent_message |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
921 |
else |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
922 |
() |
41168 | 923 |
val birth = Timer.checkRealTimer timer |
41171 | 924 |
val _ = |
41211
1e2e16bc0077
no need to do a super-duper atomization if Metis fails afterwards anyway
blanchet
parents:
41209
diff
changeset
|
925 |
if debug then Output.urgent_message "Invoking SMT solver..." else () |
41209 | 926 |
val (outcome, used_facts) = |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
927 |
(case (slice, smt_filter) of |
41432
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41336
diff
changeset
|
928 |
(1, SOME head) => head |> apsnd (apsnd repair_context) |
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41336
diff
changeset
|
929 |
| _ => SMT_Solver.smt_filter_preprocess state facts i) |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
930 |
|> SMT_Solver.smt_filter_apply slice_timeout |
41239 | 931 |
|> (fn {outcome, used_facts} => (outcome, used_facts)) |
41209 | 932 |
handle exn => if Exn.is_interrupt exn then |
933 |
reraise exn |
|
934 |
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:
42060
diff
changeset
|
935 |
(ML_Compiler.exn_message exn |
41209 | 936 |
|> SMT_Failure.Other_Failure |> SOME, []) |
41168 | 937 |
val death = Timer.checkRealTimer timer |
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
938 |
val outcome0 = if is_none outcome0 then SOME outcome else outcome0 |
41168 | 939 |
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:
40471
diff
changeset
|
940 |
val too_many_facts_perhaps = |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
941 |
case outcome of |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
942 |
NONE => false |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
943 |
| SOME (SMT_Failure.Counterexample _) => false |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
944 |
| SOME SMT_Failure.Time_Out => slice_timeout <> timeout |
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
945 |
| 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:
40471
diff
changeset
|
946 |
| SOME SMT_Failure.Out_Of_Memory => true |
41211
1e2e16bc0077
no need to do a super-duper atomization if Metis fails afterwards anyway
blanchet
parents:
41209
diff
changeset
|
947 |
| SOME (SMT_Failure.Other_Failure _) => true |
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
948 |
val timeout = Time.- (timeout, Timer.checkRealTimer timer) |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
949 |
in |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
950 |
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:
40471
diff
changeset
|
951 |
num_facts > 0 andalso Time.> (timeout, Time.zeroTime) then |
41169 | 952 |
let |
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
953 |
val new_num_facts = |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
954 |
Real.ceil (Config.get ctxt smt_slice_fact_frac |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
955 |
* Real.fromInt num_facts) |
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
956 |
val _ = |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
957 |
if verbose andalso is_some outcome then |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
958 |
quote name ^ " invoked with " ^ string_of_int num_facts ^ |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
959 |
" fact" ^ plural_s num_facts ^ ": " ^ |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
960 |
string_for_failure (failure_from_smt_failure (the outcome)) ^ |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
961 |
" 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:
42623
diff
changeset
|
962 |
plural_s new_num_facts ^ "..." |
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
963 |
|> Output.urgent_message |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
964 |
else |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
965 |
() |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
966 |
in |
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
967 |
facts |> take new_num_facts |
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
968 |
|> do_slice timeout (slice + 1) outcome0 time_so_far |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
969 |
end |
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
970 |
else |
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
971 |
{outcome = if is_none outcome then NONE else the outcome0, |
41168 | 972 |
used_facts = used_facts, |
973 |
run_time_in_msecs = SOME (Time.toMilliseconds time_so_far)} |
|
40409
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents:
40370
diff
changeset
|
974 |
end |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
975 |
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:
40370
diff
changeset
|
976 |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
977 |
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:
43006
diff
changeset
|
978 |
minimize_command |
5f8d74d3b297
added syntax for specifying Metis timeout (currently used only by SMT solvers)
blanchet
parents:
43006
diff
changeset
|
979 |
({state, subgoal, subgoal_count, facts, smt_filter, ...} |
5f8d74d3b297
added syntax for specifying Metis timeout (currently used only by SMT solvers)
blanchet
parents:
43006
diff
changeset
|
980 |
: prover_problem) = |
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
981 |
let |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
982 |
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:
41242
diff
changeset
|
983 |
val num_facts = length facts |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
984 |
val facts = facts ~~ (0 upto num_facts - 1) |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
985 |
|> map (smt_weighted_fact ctxt num_facts) |
40181 | 986 |
val {outcome, used_facts, run_time_in_msecs} = |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
987 |
smt_filter_loop ctxt name params state subgoal smt_filter facts |
43037 | 988 |
val (used_facts, used_ths) = used_facts |> ListPair.unzip |
41222 | 989 |
val outcome = outcome |> Option.map failure_from_smt_failure |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
990 |
val (preplay, message) = |
40184
91b4b73dbafb
proper error handling for SMT solvers in Sledgehammer
blanchet
parents:
40181
diff
changeset
|
991 |
case outcome of |
91b4b73dbafb
proper error handling for SMT solvers in Sledgehammer
blanchet
parents:
40181
diff
changeset
|
992 |
NONE => |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
993 |
(fn () => |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
994 |
let |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
995 |
fun smt_settings () = |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
996 |
if name = SMT_Solver.solver_name_of ctxt then "" |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
997 |
else "smt_solver = " ^ maybe_quote name |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
998 |
in |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
999 |
case play_one_line_proof debug preplay_timeout used_ths state |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1000 |
subgoal [Metis, MetisFT] of |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1001 |
p as Played _ => p |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1002 |
| _ => Trust_Playable (SMT (smt_settings ()), NONE) |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1003 |
end, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1004 |
fn preplay => |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1005 |
let |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1006 |
val one_line_params = |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1007 |
(preplay, proof_banner mode name, used_facts, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1008 |
choose_minimize_command minimize_command name preplay, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1009 |
subgoal, subgoal_count) |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1010 |
in |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1011 |
one_line_proof_text one_line_params ^ |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1012 |
(if verbose then |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1013 |
"\nSMT solver real CPU time: " ^ |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1014 |
string_from_time (Time.fromMilliseconds |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1015 |
(the run_time_in_msecs)) ^ "." |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1016 |
else |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1017 |
"") |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1018 |
end) |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1019 |
| SOME failure => (K Failed_to_Play, fn _ => string_for_failure failure) |
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1020 |
in |
43037 | 1021 |
{outcome = outcome, used_facts = used_facts, |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1022 |
run_time_in_msecs = run_time_in_msecs, preplay = preplay, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1023 |
message = message} |
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1024 |
end |
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1025 |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1026 |
fun run_metis mode name ({debug, timeout, ...} : params) minimize_command |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1027 |
({state, subgoal, subgoal_count, facts, ...} : prover_problem) = |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1028 |
let |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1029 |
val reconstructor = if name = Metis_Tactics.metisN then Metis |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1030 |
else if name = Metis_Tactics.metisFT_N then MetisFT |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1031 |
else raise Fail ("unknown Metis version: " ^ quote name) |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1032 |
val (used_facts, used_ths) = |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1033 |
facts |> map untranslated_fact |> ListPair.unzip |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1034 |
in |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1035 |
case play_one_line_proof debug timeout used_ths state subgoal |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1036 |
[reconstructor] of |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1037 |
play as Played (_, time) => |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1038 |
{outcome = NONE, used_facts = used_facts, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1039 |
run_time_in_msecs = SOME (Time.toMilliseconds time), |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1040 |
preplay = K play, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1041 |
message = fn play => |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1042 |
let |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1043 |
val one_line_params = |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1044 |
(play, proof_banner mode name, used_facts, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1045 |
minimize_command name, subgoal, subgoal_count) |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1046 |
in one_line_proof_text one_line_params end} |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1047 |
| play => |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1048 |
let val failure = if play = Failed_to_Play then GaveUp else TimedOut in |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1049 |
{outcome = SOME failure, used_facts = [], run_time_in_msecs = NONE, |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1050 |
preplay = K play, message = fn _ => string_for_failure failure} |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1051 |
end |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1052 |
end |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1053 |
|
43021 | 1054 |
fun get_prover ctxt mode name = |
42361 | 1055 |
let val thy = Proof_Context.theory_of ctxt in |
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1056 |
if is_metis_prover name then run_metis mode name |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1057 |
else if is_atp thy name then run_atp mode name (get_atp thy name) |
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1058 |
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:
43051
diff
changeset
|
1059 |
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:
40723
diff
changeset
|
1060 |
end |
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1061 |
|
28582 | 1062 |
end; |