author | wenzelm |
Sat, 16 Apr 2011 16:15:37 +0200 | |
changeset 42361 | 23f352990944 |
parent 42180 | a6c141925a8a |
child 42443 | 724e612ba248 |
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_run.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 |
|
38021
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents:
38020
diff
changeset
|
6 |
Sledgehammer's heart. |
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_RUN = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
10 |
sig |
38988 | 11 |
type relevance_override = Sledgehammer_Filter.relevance_override |
40068 | 12 |
type minimize_command = Sledgehammer_ATP_Reconstruct.minimize_command |
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
13 |
type params = Sledgehammer_Provers.params |
41263
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
14 |
type prover = Sledgehammer_Provers.prover |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
15 |
|
41335
66edbd0f7a2e
added "smt_triggers" option to experiment with triggers in Sledgehammer;
blanchet
parents:
41316
diff
changeset
|
16 |
val auto_minimize_min_facts : int Unsynchronized.ref |
41263
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
17 |
val get_minimizing_prover : Proof.context -> bool -> string -> prover |
38044 | 18 |
val run_sledgehammer : |
39318 | 19 |
params -> bool -> int -> relevance_override -> (string -> minimize_command) |
20 |
-> Proof.state -> bool * Proof.state |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
21 |
end; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
22 |
|
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
23 |
structure Sledgehammer_Run : SLEDGEHAMMER_RUN = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
24 |
struct |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
25 |
|
38023 | 26 |
open Sledgehammer_Util |
38988 | 27 |
open Sledgehammer_Filter |
40068 | 28 |
open Sledgehammer_ATP_Translate |
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
29 |
open Sledgehammer_Provers |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
30 |
open Sledgehammer_Minimize |
40072
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents:
40071
diff
changeset
|
31 |
|
41208
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents:
41180
diff
changeset
|
32 |
fun prover_description ctxt ({verbose, blocking, ...} : params) name num_facts i |
41089 | 33 |
n goal = |
34 |
quote name ^ |
|
35 |
(if verbose then |
|
36 |
" with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts |
|
37 |
else |
|
38 |
"") ^ |
|
41743 | 39 |
" on " ^ (if n = 1 then "goal" else "subgoal " ^ string_of_int i) ^ |
41089 | 40 |
(if blocking then |
41743 | 41 |
"." |
41089 | 42 |
else |
41743 | 43 |
":\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i))) |
41089 | 44 |
|
41335
66edbd0f7a2e
added "smt_triggers" option to experiment with triggers in Sledgehammer;
blanchet
parents:
41316
diff
changeset
|
45 |
val auto_minimize_min_facts = Unsynchronized.ref (!binary_min_facts) |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
46 |
|
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
47 |
fun get_minimizing_prover ctxt auto name |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
48 |
(params as {debug, verbose, explicit_apply, ...}) minimize_command |
41263
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
49 |
(problem as {state, subgoal, subgoal_count, facts, ...}) = |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
50 |
get_prover ctxt auto name params minimize_command problem |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
51 |
|> (fn result as {outcome, used_facts, run_time_in_msecs, message} => |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
52 |
if is_some outcome then |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
53 |
result |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
54 |
else |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
55 |
let |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
56 |
val (used_facts, message) = |
41335
66edbd0f7a2e
added "smt_triggers" option to experiment with triggers in Sledgehammer;
blanchet
parents:
41316
diff
changeset
|
57 |
if length used_facts >= !auto_minimize_min_facts then |
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
58 |
minimize_facts name params (SOME explicit_apply) (not verbose) |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
59 |
subgoal subgoal_count state |
41263
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
60 |
(filter_used_facts used_facts |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
61 |
(map (apsnd single o untranslated_fact) facts)) |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
62 |
|>> Option.map (map fst) |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
63 |
else |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
64 |
(SOME used_facts, message) |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
65 |
in |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
66 |
case used_facts of |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
67 |
SOME used_facts => |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
68 |
(if debug andalso not (null used_facts) then |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
69 |
facts ~~ (0 upto length facts - 1) |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
70 |
|> map (fn (fact, j) => |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
71 |
fact |> untranslated_fact |> apsnd (K j)) |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
72 |
|> filter_used_facts used_facts |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
73 |
|> map (fn ((name, _), j) => name ^ "@" ^ string_of_int j) |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
74 |
|> commas |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
75 |
|> enclose ("Fact" ^ plural_s (length facts) ^ " in " ^ |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
76 |
quote name ^ " proof (of " ^ |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
77 |
string_of_int (length facts) ^ "): ") "." |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
78 |
|> Output.urgent_message |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
79 |
else |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
80 |
(); |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
81 |
{outcome = NONE, used_facts = used_facts, |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
82 |
run_time_in_msecs = run_time_in_msecs, message = message}) |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
83 |
| NONE => result |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
84 |
end) |
41262
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
85 |
|
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
86 |
fun launch_prover |
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
87 |
(params as {debug, blocking, max_relevant, timeout, expect, ...}) |
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
88 |
auto minimize_command only |
41741 | 89 |
{state, goal, subgoal, subgoal_count, facts, smt_filter} name = |
41089 | 90 |
let |
91 |
val ctxt = Proof.context_of state |
|
92 |
val birth_time = Time.now () |
|
93 |
val death_time = Time.+ (birth_time, timeout) |
|
94 |
val max_relevant = |
|
95 |
the_default (default_max_relevant_for_prover ctxt name) max_relevant |
|
96 |
val num_facts = length facts |> not only ? Integer.min max_relevant |
|
97 |
val desc = |
|
98 |
prover_description ctxt params name num_facts subgoal subgoal_count goal |
|
99 |
val problem = |
|
100 |
{state = state, goal = goal, subgoal = subgoal, |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
101 |
subgoal_count = subgoal_count, facts = take num_facts facts, |
41741 | 102 |
smt_filter = smt_filter} |
41255
a80024d7b71b
added debugging option to find out how good the relevance filter was at identifying relevant facts
blanchet
parents:
41245
diff
changeset
|
103 |
fun really_go () = |
41263
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
104 |
problem |
4cac389c005f
renamed function to run prover with minimizer and changed signature to clarify its semantics and make it a drop in replacement for "get_prover"
blanchet
parents:
41262
diff
changeset
|
105 |
|> get_minimizing_prover ctxt auto name params (minimize_command name) |
41262
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
106 |
|> (fn {outcome, message, ...} => |
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
107 |
(if is_some outcome then "none" else "some" (* sic *), message)) |
41089 | 108 |
fun go () = |
109 |
let |
|
110 |
val (outcome_code, message) = |
|
111 |
if debug then |
|
112 |
really_go () |
|
113 |
else |
|
114 |
(really_go () |
|
115 |
handle ERROR message => ("unknown", "Error: " ^ message ^ "\n") |
|
116 |
| exn => |
|
117 |
if Exn.is_interrupt exn then |
|
118 |
reraise exn |
|
119 |
else |
|
120 |
("unknown", "Internal error:\n" ^ |
|
121 |
ML_Compiler.exn_message exn ^ "\n")) |
|
122 |
val _ = |
|
41142
43e2b051339c
weaken the "expect" flag so that it doesn't trigger errors if a prover is not installed
blanchet
parents:
41138
diff
changeset
|
123 |
(* The "expect" argument is deliberately ignored if the prover is |
43e2b051339c
weaken the "expect" flag so that it doesn't trigger errors if a prover is not installed
blanchet
parents:
41138
diff
changeset
|
124 |
missing so that the "Metis_Examples" can be processed on any |
43e2b051339c
weaken the "expect" flag so that it doesn't trigger errors if a prover is not installed
blanchet
parents:
41138
diff
changeset
|
125 |
machine. *) |
43e2b051339c
weaken the "expect" flag so that it doesn't trigger errors if a prover is not installed
blanchet
parents:
41138
diff
changeset
|
126 |
if expect = "" orelse outcome_code = expect orelse |
43e2b051339c
weaken the "expect" flag so that it doesn't trigger errors if a prover is not installed
blanchet
parents:
41138
diff
changeset
|
127 |
not (is_prover_installed ctxt name) then |
41089 | 128 |
() |
129 |
else if blocking then |
|
130 |
error ("Unexpected outcome: " ^ quote outcome_code ^ ".") |
|
131 |
else |
|
132 |
warning ("Unexpected outcome: " ^ quote outcome_code ^ "."); |
|
133 |
in (outcome_code = "some", message) end |
|
134 |
in |
|
135 |
if auto then |
|
136 |
let val (success, message) = TimeLimit.timeLimit timeout go () in |
|
137 |
(success, state |> success ? Proof.goal_message (fn () => |
|
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
138 |
Pretty.chunks [Pretty.str "", |
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
139 |
Pretty.mark Markup.hilite (Pretty.str message)])) |
41089 | 140 |
end |
141 |
else if blocking then |
|
142 |
let val (success, message) = TimeLimit.timeLimit timeout go () in |
|
143 |
List.app Output.urgent_message |
|
144 |
(Async_Manager.break_into_chunks [desc ^ "\n" ^ message]); |
|
145 |
(success, state) |
|
146 |
end |
|
147 |
else |
|
148 |
(Async_Manager.launch das_Tool birth_time death_time desc (snd o go); |
|
149 |
(false, state)) |
|
150 |
end |
|
151 |
||
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
152 |
fun class_of_smt_solver ctxt name = |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
153 |
ctxt |> select_smt_solver name |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
154 |
|> SMT_Config.solver_class_of |> SMT_Utils.string_of_class |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
155 |
|
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
156 |
(* Makes backtraces more transparent and might be more efficient as well. *) |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
157 |
fun smart_par_list_map _ [] = [] |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
158 |
| smart_par_list_map f [x] = [f x] |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
159 |
| smart_par_list_map f xs = Par_List.map f xs |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
160 |
|
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
161 |
fun dest_SMT_Weighted_Fact (SMT_Weighted_Fact p) = p |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
162 |
| dest_SMT_Weighted_Fact _ = raise Fail "dest_SMT_Weighted_Fact" |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
163 |
|
40698 | 164 |
(* FUDGE *) |
165 |
val auto_max_relevant_divisor = 2 |
|
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
166 |
|
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
167 |
fun run_sledgehammer (params as {debug, blocking, provers, monomorphize, |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
168 |
type_sys, relevance_thresholds, max_relevant, |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
169 |
timeout, ...}) |
39366
f58fbb959826
handle relevance filter corner cases more gracefully;
blanchet
parents:
39364
diff
changeset
|
170 |
auto i (relevance_override as {only, ...}) minimize_command |
f58fbb959826
handle relevance filter corner cases more gracefully;
blanchet
parents:
39364
diff
changeset
|
171 |
state = |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
172 |
if null provers then |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
173 |
error "No prover is set." |
39318 | 174 |
else case subgoal_count state of |
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
40114
diff
changeset
|
175 |
0 => (Output.urgent_message "No subgoal!"; (false, state)) |
39318 | 176 |
| n => |
177 |
let |
|
39364 | 178 |
val _ = Proof.assert_backward state |
41773 | 179 |
val print = if auto then K () else Output.urgent_message |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
180 |
val state = |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
181 |
state |> Proof.map_context (Config.put SMT_Config.verbose debug) |
40200 | 182 |
val ctxt = Proof.context_of state |
42361 | 183 |
val thy = Proof_Context.theory_of ctxt |
40200 | 184 |
val {facts = chained_ths, goal, ...} = Proof.goal state |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
185 |
val (_, hyp_ts, concl_t) = strip_subgoal goal i |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41091
diff
changeset
|
186 |
val no_dangerous_types = types_dangerous_types type_sys |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
187 |
val _ = () |> not blocking ? kill_provers |
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41432
diff
changeset
|
188 |
val _ = case find_first (not o is_prover_supported ctxt) provers of |
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
189 |
SOME name => error ("No such prover: " ^ name ^ ".") |
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
190 |
| NONE => () |
41773 | 191 |
val _ = print "Sledgehammering..." |
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
192 |
val (smts, atps) = provers |> List.partition (is_smt_prover ctxt) |
41741 | 193 |
fun launch_provers state get_facts translate maybe_smt_filter provers = |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
194 |
let |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
195 |
val facts = get_facts () |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
196 |
val num_facts = length facts |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
197 |
val facts = facts ~~ (0 upto num_facts - 1) |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
198 |
|> map (translate num_facts) |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
199 |
val problem = |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
200 |
{state = state, goal = goal, subgoal = i, subgoal_count = n, |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
201 |
facts = facts, |
41741 | 202 |
smt_filter = maybe_smt_filter |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
203 |
(fn () => map_filter (try dest_SMT_Weighted_Fact) facts) i} |
41262
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
204 |
val launch = launch_prover params auto minimize_command only |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
205 |
in |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
206 |
if auto then |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
207 |
fold (fn prover => fn (true, state) => (true, state) |
41262
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
208 |
| (false, _) => launch problem prover) |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
209 |
provers (false, state) |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
210 |
else |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
211 |
provers |
41262
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
212 |
|> (if blocking then smart_par_list_map else map) (launch problem) |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
213 |
|> exists fst |> rpair state |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
214 |
end |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
215 |
fun get_facts label no_dangerous_types relevance_fudge provers = |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
216 |
let |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
217 |
val max_max_relevant = |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
218 |
case max_relevant of |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
219 |
SOME n => n |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
220 |
| NONE => |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
221 |
0 |> fold (Integer.max o default_max_relevant_for_prover ctxt) |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
222 |
provers |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
223 |
|> auto ? (fn n => n div auto_max_relevant_divisor) |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
224 |
val is_built_in_const = |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
225 |
is_built_in_const_for_prover ctxt (hd provers) |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
226 |
in |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
227 |
relevant_facts ctxt no_dangerous_types relevance_thresholds |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
228 |
max_max_relevant is_built_in_const relevance_fudge |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
229 |
relevance_override chained_ths hyp_ts concl_t |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
230 |
|> tap (fn facts => |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
231 |
if debug then |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
232 |
label ^ plural_s (length provers) ^ ": " ^ |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
233 |
(if null facts then |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
234 |
"Found no relevant facts." |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
235 |
else |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
236 |
"Including (up to) " ^ string_of_int (length facts) ^ |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
237 |
" relevant fact" ^ plural_s (length facts) ^ ":\n" ^ |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
238 |
(facts |> map (fst o fst) |> space_implode " ") ^ ".") |
41773 | 239 |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
240 |
else |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
241 |
()) |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
242 |
end |
41746
e590971528b2
run all provers in blocking mode, even if a proof was already found -- this behavior is less confusing to the user
blanchet
parents:
41743
diff
changeset
|
243 |
fun launch_atps accum = |
e590971528b2
run all provers in blocking mode, even if a proof was already found -- this behavior is less confusing to the user
blanchet
parents:
41743
diff
changeset
|
244 |
if null atps then |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
245 |
accum |
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
246 |
else |
41262
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
247 |
launch_provers state |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
248 |
(get_facts "ATP" no_dangerous_types atp_relevance_fudge o K atps) |
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
249 |
(if monomorphize then |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
250 |
K (Untranslated_Fact o fst) |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
251 |
else |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41990
diff
changeset
|
252 |
ATP_Translated_Fact oo K (translate_atp_fact ctxt false o fst)) |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
253 |
(K (K NONE)) atps |
41746
e590971528b2
run all provers in blocking mode, even if a proof was already found -- this behavior is less confusing to the user
blanchet
parents:
41743
diff
changeset
|
254 |
fun launch_smts accum = |
e590971528b2
run all provers in blocking mode, even if a proof was already found -- this behavior is less confusing to the user
blanchet
parents:
41743
diff
changeset
|
255 |
if null smts then |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
256 |
accum |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
257 |
else |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
258 |
let |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
259 |
val facts = get_facts "SMT solver" true smt_relevance_fudge smts |
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41255
diff
changeset
|
260 |
val weight = SMT_Weighted_Fact oo weight_smt_fact thy |
41741 | 261 |
fun smt_filter facts = |
41788
adfd365c7ea4
added a timeout around SMT preprocessing (notably monomorphization)
blanchet
parents:
41773
diff
changeset
|
262 |
(if debug then curry (op o) SOME |
adfd365c7ea4
added a timeout around SMT preprocessing (notably monomorphization)
blanchet
parents:
41773
diff
changeset
|
263 |
else TimeLimit.timeLimit timeout o try) |
41432
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41335
diff
changeset
|
264 |
(SMT_Solver.smt_filter_preprocess state (facts ())) |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
265 |
in |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
266 |
smts |> map (`(class_of_smt_solver ctxt)) |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
267 |
|> AList.group (op =) |
41741 | 268 |
|> map (launch_provers state (K facts) weight smt_filter o snd) |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
269 |
|> exists fst |> rpair state |
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
270 |
end |
41262
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
271 |
fun launch_atps_and_smt_solvers () = |
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
272 |
[launch_atps, launch_smts] |
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41208
diff
changeset
|
273 |
|> smart_par_list_map (fn f => f (false, state) |> K ()) |
41773 | 274 |
handle ERROR msg => (print ("Error: " ^ msg); error msg) |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
275 |
in |
40065 | 276 |
(false, state) |
41262
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
277 |
|> (if blocking then launch_atps #> not auto ? launch_smts |
095ecb0c687f
factored out running a prover with (optionally) an implicit minimizer phrase
blanchet
parents:
41260
diff
changeset
|
278 |
else (fn p => Future.fork (tap launch_atps_and_smt_solvers) |> K p)) |
41773 | 279 |
handle TimeLimit.TimeOut => |
280 |
(print "Sledgehammer ran out of time."; (false, state)) |
|
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
281 |
end |
38044 | 282 |
|
28582 | 283 |
end; |