author | desharna |
Wed, 30 Mar 2022 10:37:38 +0200 | |
changeset 75373 | 48736d743e8c |
parent 75372 | 4c8d1ef258d3 |
child 75374 | 6e8ca4959334 |
permissions | -rw-r--r-- |
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer.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 |
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
9 |
signature SLEDGEHAMMER = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
10 |
sig |
60612
79d71bfea310
removed chained facts from preplaying -- and careful about extra chained facts when removing 'proof -' and 'qed' from one-line Isar proofs
blanchet
parents:
60549
diff
changeset
|
11 |
type stature = ATP_Problem_Generate.stature |
51008 | 12 |
type fact = Sledgehammer_Fact.fact |
48292 | 13 |
type fact_override = Sledgehammer_Fact.fact_override |
57755 | 14 |
type proof_method = Sledgehammer_Proof_Methods.proof_method |
15 |
type play_outcome = Sledgehammer_Proof_Methods.play_outcome |
|
55201 | 16 |
type mode = Sledgehammer_Prover.mode |
17 |
type params = Sledgehammer_Prover.params |
|
74951
0b6f795d3b78
proper filtering inf induction rules in Mirabelle
desharna
parents:
74950
diff
changeset
|
18 |
type induction_rules = Sledgehammer_Prover.induction_rules |
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
19 |
type prover_problem = Sledgehammer_Prover.prover_problem |
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
20 |
type prover_result = Sledgehammer_Prover.prover_result |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
21 |
|
75372 | 22 |
type preplay_result = proof_method * play_outcome * (string * stature) list |
23 |
||
75020
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
24 |
datatype sledgehammer_outcome = |
75372 | 25 |
SH_Some of prover_result * preplay_result list |
75020
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
26 |
| SH_Unknown |
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
27 |
| SH_Timeout |
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
28 |
| SH_None |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
29 |
|
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
30 |
val short_string_of_sledgehammer_outcome : sledgehammer_outcome -> string |
51010 | 31 |
val string_of_factss : (string * fact list) list -> string |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54799
diff
changeset
|
32 |
val run_sledgehammer : params -> mode -> (string -> unit) option -> int -> fact_override -> |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
33 |
Proof.state -> bool * (sledgehammer_outcome * string) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
34 |
end; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
35 |
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
36 |
structure Sledgehammer : SLEDGEHAMMER = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
37 |
struct |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
38 |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
39 |
open ATP_Util |
75029 | 40 |
open ATP_Problem |
55212 | 41 |
open ATP_Proof |
46320 | 42 |
open ATP_Problem_Generate |
38023 | 43 |
open Sledgehammer_Util |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
47904
diff
changeset
|
44 |
open Sledgehammer_Fact |
55287 | 45 |
open Sledgehammer_Proof_Methods |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
46 |
open Sledgehammer_Isar_Proof |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
47 |
open Sledgehammer_Isar_Preplay |
57740 | 48 |
open Sledgehammer_Isar_Minimize |
75029 | 49 |
open Sledgehammer_ATP_Systems |
55201 | 50 |
open Sledgehammer_Prover |
55212 | 51 |
open Sledgehammer_Prover_ATP |
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
52 |
open Sledgehammer_Prover_Minimize |
48381 | 53 |
open Sledgehammer_MaSh |
40072
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents:
40071
diff
changeset
|
54 |
|
75372 | 55 |
type preplay_result = proof_method * play_outcome * (string * stature) list |
56 |
||
75020
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
57 |
datatype sledgehammer_outcome = |
75372 | 58 |
SH_Some of prover_result * preplay_result list |
75020
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
59 |
| SH_Unknown |
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
60 |
| SH_Timeout |
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
61 |
| SH_None |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
62 |
|
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
63 |
fun short_string_of_sledgehammer_outcome (SH_Some _) = "some" |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
64 |
| short_string_of_sledgehammer_outcome SH_Unknown = "unknown" |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
65 |
| short_string_of_sledgehammer_outcome SH_Timeout = "timeout" |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
66 |
| short_string_of_sledgehammer_outcome SH_None = "none" |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
67 |
|
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
68 |
fun alternative f (SOME x) (SOME y) = SOME (f (x, y)) |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
69 |
| alternative _ (x as SOME _) NONE = x |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
70 |
| alternative _ NONE (y as SOME _) = y |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
71 |
| alternative _ NONE NONE = NONE |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
72 |
|
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
73 |
fun max_outcome outcomes = |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
74 |
let |
75046 | 75 |
val some = find_first (fn (SH_Some _, _) => true | _ => false) outcomes |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
76 |
val unknown = find_first (fn (SH_Unknown, _) => true | _ => false) outcomes |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
77 |
val timeout = find_first (fn (SH_Timeout, _) => true | _ => false) outcomes |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
78 |
val none = find_first (fn (SH_None, _) => true | _ => false) outcomes |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
79 |
in |
75046 | 80 |
some |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
81 |
|> alternative snd unknown |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
82 |
|> alternative snd timeout |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
83 |
|> alternative snd none |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
84 |
|> the_default (SH_Unknown, "") |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
85 |
end |
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
86 |
|
75372 | 87 |
fun play_one_line_proofs minimize timeout used_facts state i methss = |
63311
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
88 |
(if timeout = Time.zeroTime then |
75372 | 89 |
[] |
63311
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
90 |
else |
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
91 |
let |
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
92 |
val ctxt = Proof.context_of state |
75372 | 93 |
val used_facts = filter_out (fn (_, (sc, _)) => sc = Chained) used_facts |
94 |
val fact_names = map fst used_facts |
|
63311
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
95 |
val {facts = chained, goal, ...} = Proof.goal state |
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
96 |
val goal_t = Logic.get_goal (Thm.prop_of goal) i |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
97 |
|
75372 | 98 |
fun try_methss ress [] = ress |
63311
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
99 |
| try_methss ress (meths :: methss) = |
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
100 |
let |
75372 | 101 |
fun mk_step meths = |
72584 | 102 |
Prove { |
103 |
qualifiers = [], |
|
104 |
obtains = [], |
|
105 |
label = ("", 0), |
|
106 |
goal = goal_t, |
|
107 |
subproofs = [], |
|
108 |
facts = ([], fact_names), |
|
109 |
proof_methods = meths, |
|
110 |
comment = ""} |
|
75372 | 111 |
val ress' = |
112 |
preplay_isar_step ctxt chained timeout [] (mk_step meths) |
|
113 |
|> map (fn result as (meth, play_outcome) => |
|
114 |
(case (minimize, play_outcome) of |
|
115 |
(true, Played time) => |
|
116 |
let |
|
117 |
val (time', used_names') = |
|
118 |
minimized_isar_step ctxt chained time (mk_step [meth]) |
|
119 |
||> (facts_of_isar_step #> snd) |
|
120 |
val used_facts' = filter (member (op =) used_names' o fst) used_facts |
|
121 |
in |
|
122 |
(meth, Played time', used_facts') |
|
123 |
end |
|
124 |
| _ => (meth, play_outcome, used_facts))) |
|
125 |
val any_succeeded = exists (fn (_, Played _, _) => true | _ => false) ress' |
|
63311
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
126 |
in |
75372 | 127 |
try_methss (ress' @ ress) (if any_succeeded then [] else methss) |
63311
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
128 |
end |
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
129 |
in |
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
130 |
try_methss [] methss |
540cfb14a751
be more careful before filtering out chained facts in Sledgehammer
blanchet
parents:
63097
diff
changeset
|
131 |
end) |
75372 | 132 |
|> map (fn (meth, play_outcome, used_facts) => (meth, play_outcome, filter_out (fn (_, (sc, _)) => sc = Chained) used_facts)) |
133 |
|> sort (play_outcome_ord o apply2 (fn (_, play_outcome, _) => play_outcome)) |
|
134 |
||
135 |
fun select_one_line_proof used_facts preferred_meth preplay_results = |
|
136 |
(case preplay_results of |
|
137 |
[] => (used_facts, (preferred_meth, Play_Timed_Out Time.zeroTime)) |
|
138 |
| (best_meth, best_outcome, best_used_facts) :: results' => |
|
139 |
let |
|
140 |
val (prefered_outcome, prefered_used_facts) = |
|
141 |
(case find_first (fn (meth, _, _) => meth = preferred_meth) preplay_results of |
|
142 |
NONE => (Play_Timed_Out Time.zeroTime, used_facts) |
|
143 |
| SOME (_, prefered_outcome, prefered_used_facts) => |
|
144 |
(prefered_outcome, prefered_used_facts)) |
|
145 |
in |
|
146 |
(case (prefered_outcome, best_outcome) of |
|
147 |
(* If prefered_meth succeeded, use it irrespective of other preplay results *) |
|
148 |
(Played _, _) => (prefered_used_facts, (preferred_meth, prefered_outcome)) |
|
149 |
(* If prefered_meth did not succeed but best method did, use best method *) |
|
150 |
| (_, Played _) => (best_used_facts, (best_meth, best_outcome)) |
|
151 |
(* If neither succeeded, use preferred_meth *) |
|
152 |
| (_, _) => (prefered_used_facts, (preferred_meth, prefered_outcome))) |
|
153 |
end) |
|
154 |
|> apfst (filter_out (fn (_, (sc, _)) => sc = Chained)) |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
155 |
|
75035 | 156 |
fun launch_prover (params as {verbose, spy, slices, timeout, ...}) mode learn |
75033 | 157 |
(problem as {state, subgoal, factss, ...} : prover_problem) |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
158 |
(slice as ((slice_size, num_facts, fact_filter), _)) name = |
41089 | 159 |
let |
160 |
val ctxt = Proof.context_of state |
|
53800 | 161 |
|
73975
8d93f9ca6518
revisited ac28714b7478: more faithful preplaying with chained facts
blanchet
parents:
73940
diff
changeset
|
162 |
val _ = spying spy (fn () => (state, subgoal, name, "Launched")) |
53800 | 163 |
|
75033 | 164 |
val _ = |
165 |
if verbose then |
|
166 |
writeln (name ^ " with " ^ string_of_int num_facts ^ " " ^ fact_filter ^ " fact" ^ |
|
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
167 |
plural_s num_facts ^ " for " ^ string_of_time (slice_timeout slice_size slices timeout) ^ |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
168 |
"...") |
75033 | 169 |
else |
170 |
() |
|
171 |
||
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51008
diff
changeset
|
172 |
fun print_used_facts used_facts used_from = |
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51008
diff
changeset
|
173 |
tag_list 1 used_from |
51005
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents:
51004
diff
changeset
|
174 |
|> map (fn (j, fact) => fact |> apsnd (K j)) |
48798 | 175 |
|> filter_used_facts false used_facts |
48394
82fc8c956cdc
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
blanchet
parents:
48384
diff
changeset
|
176 |
|> map (fn ((name, _), j) => name ^ "@" ^ string_of_int j) |
82fc8c956cdc
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
blanchet
parents:
48384
diff
changeset
|
177 |
|> commas |
75025 | 178 |
|> prefix ("Facts in " ^ name ^ " proof: ") |
58843 | 179 |
|> writeln |
53800 | 180 |
|
54062 | 181 |
fun spying_str_of_res ({outcome = NONE, used_facts, used_from, ...} : prover_result) = |
182 |
let |
|
183 |
val num_used_facts = length used_facts |
|
54063 | 184 |
|
185 |
fun find_indices facts = |
|
186 |
tag_list 1 facts |
|
54062 | 187 |
|> map (fn (j, fact) => fact |> apsnd (K j)) |
188 |
|> filter_used_facts false used_facts |
|
54773 | 189 |
|> distinct (eq_fst (op =)) |
54062 | 190 |
|> map (prefix "@" o string_of_int o snd) |
54063 | 191 |
|
192 |
fun filter_info (fact_filter, facts) = |
|
193 |
let |
|
194 |
val indices = find_indices facts |
|
75033 | 195 |
(* "Int.max" is there for robustness *) |
54773 | 196 |
val unknowns = replicate (Int.max (0, num_used_facts - length indices)) "?" |
197 |
in |
|
198 |
(commas (indices @ unknowns), fact_filter) |
|
199 |
end |
|
54063 | 200 |
|
201 |
val filter_infos = |
|
75025 | 202 |
map filter_info (("actual", used_from) :: factss) |
54063 | 203 |
|> AList.group (op =) |
204 |
|> map (fn (indices, fact_filters) => commas fact_filters ^ ": " ^ indices) |
|
54062 | 205 |
in |
75025 | 206 |
"Success: Found proof with " ^ string_of_int num_used_facts ^ " fact" ^ |
207 |
plural_s num_used_facts ^ |
|
54063 | 208 |
(if num_used_facts = 0 then "" else ": " ^ commas filter_infos) |
53800 | 209 |
end |
210 |
| spying_str_of_res {outcome = SOME failure, ...} = |
|
54062 | 211 |
"Failure: " ^ string_of_atp_failure failure |
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
212 |
in |
75033 | 213 |
get_minimizing_prover ctxt mode learn name params problem slice |
214 |
|> verbose ? tap (fn {outcome = NONE, used_facts as _ :: _, used_from, ...} => |
|
215 |
print_used_facts used_facts used_from |
|
216 |
| _ => ()) |
|
217 |
|> spy ? tap (fn res => spying spy (fn () => (state, subgoal, name, spying_str_of_res res))) |
|
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
218 |
end |
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
219 |
|
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
220 |
fun preplay_prover_result ({ minimize, preplay_timeout, ...} : params) state subgoal |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
221 |
(result as {outcome, used_facts, preferred_methss, message, ...} : prover_result) = |
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
222 |
let |
75372 | 223 |
val (output, chosen_preplay_outcome) = |
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
224 |
if outcome = SOME ATP_Proof.TimedOut then |
75372 | 225 |
(SH_Timeout, select_one_line_proof used_facts (fst preferred_methss) []) |
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
226 |
else if is_some outcome then |
75372 | 227 |
(SH_None, select_one_line_proof used_facts (fst preferred_methss) []) |
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
228 |
else |
75372 | 229 |
let |
230 |
val preplay_results = |
|
231 |
play_one_line_proofs minimize preplay_timeout used_facts state subgoal |
|
232 |
(snd preferred_methss) |
|
233 |
val chosen_preplay_outcome = |
|
234 |
select_one_line_proof used_facts (fst preferred_methss) preplay_results |
|
235 |
in |
|
236 |
(SH_Some (result, preplay_results), chosen_preplay_outcome) |
|
237 |
end |
|
238 |
fun output_message () = message (fn () => chosen_preplay_outcome) |
|
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
239 |
in |
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
240 |
(output, output_message) |
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
241 |
end |
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
242 |
|
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
243 |
fun check_expected_outcome ctxt prover_name expect outcome = |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
244 |
let |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
245 |
val outcome_code = short_string_of_sledgehammer_outcome outcome |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
246 |
in |
75027
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
247 |
(* The "expect" argument is deliberately ignored if the prover is missing so that |
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
248 |
"Metis_Examples" can be processed on any machine. *) |
75373
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
249 |
if expect = "" orelse not (is_prover_installed ctxt prover_name) then |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
250 |
() |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
251 |
else |
75373
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
252 |
(case (expect, outcome) of |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
253 |
("some", SH_Some _) => () |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
254 |
| ("some_preplayed", SH_Some (_, preplay_results)) => |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
255 |
if exists (fn (_, Played _, _) => true | _ => false) preplay_results then |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
256 |
() |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
257 |
else |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
258 |
error ("Unexpected outcome: the external prover found a some proof but preplay failed") |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
259 |
| ("unknown", SH_Unknown) => () |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
260 |
| ("timeout", SH_Timeout) => () |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
261 |
| ("none", SH_None) => () |
48736d743e8c
expanded sledgehammer's expect option with some_preplayed
desharna
parents:
75372
diff
changeset
|
262 |
| _ => error ("Unexpected outcome: " ^ quote outcome_code)) |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
263 |
end |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
264 |
|
75027
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
265 |
fun launch_prover_and_preplay (params as {debug, timeout, expect, ...}) mode writeln_result learn |
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
266 |
(problem as {state, subgoal, ...}) slice prover_name = |
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
267 |
let |
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
268 |
val ctxt = Proof.context_of state |
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
269 |
val hard_timeout = Time.scale 5.0 timeout |
53800 | 270 |
|
41255
a80024d7b71b
added debugging option to find out how good the relevance filter was at identifying relevant facts
blanchet
parents:
41245
diff
changeset
|
271 |
fun really_go () = |
75027
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
272 |
launch_prover params mode learn problem slice prover_name |
74952
ae2185967e67
exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents:
74951
diff
changeset
|
273 |
|> preplay_prover_result params state subgoal |
53800 | 274 |
|
41089 | 275 |
fun go () = |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
276 |
if debug then |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
277 |
really_go () |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
278 |
else |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
279 |
(really_go () |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
280 |
handle |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
281 |
ERROR msg => (SH_Unknown, fn () => "Error: " ^ msg ^ "\n") |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
282 |
| exn => |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
283 |
if Exn.is_interrupt exn then Exn.reraise exn |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
284 |
else (SH_Unknown, fn () => "Internal error:\n" ^ Runtime.exn_message exn ^ "\n")) |
57056 | 285 |
|
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
286 |
val (outcome, message) = Timeout.apply hard_timeout go () |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
287 |
val () = check_expected_outcome ctxt prover_name expect outcome |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
288 |
|
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
289 |
val message = message () |
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
290 |
val () = |
75040 | 291 |
if mode = Auto_Try then |
292 |
() |
|
293 |
else |
|
294 |
(case outcome of |
|
295 |
SH_Some _ => the_default writeln writeln_result (prover_name ^ ": " ^ message) |
|
296 |
| _ => ()) |
|
41089 | 297 |
in |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
298 |
(outcome, message) |
41089 | 299 |
end |
300 |
||
75034 | 301 |
fun string_of_facts filter facts = |
302 |
"Selected " ^ string_of_int (length facts) ^ " " ^ (if filter = "" then "" else filter ^ " ") ^ |
|
303 |
"fact" ^ plural_s (length facts) ^ ": " ^ (space_implode " " (map (fst o fst) facts)) |
|
51008 | 304 |
|
51010 | 305 |
fun string_of_factss factss = |
57464 | 306 |
if forall (null o snd) factss then |
63692 | 307 |
"Found no relevant facts" |
57464 | 308 |
else |
75034 | 309 |
cat_lines (map (fn (filter, facts) => string_of_facts filter facts) factss) |
51008 | 310 |
|
75029 | 311 |
val default_slice_schedule = |
75339
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
312 |
(* FUDGE (inspired by Seventeen evaluation) *) |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
313 |
[cvc4N, zipperpositionN, vampireN, veritN, eN, cvc4N, zipperpositionN, cvc4N, vampireN, cvc4N, |
75339
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
314 |
cvc4N, vampireN, cvc4N, iproverN, zipperpositionN, vampireN, vampireN, zipperpositionN, z3N, |
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
315 |
cvc4N, vampireN, iproverN, vampireN, zipperpositionN, z3N, z3N, cvc4N, cvc4N] |
75029 | 316 |
|
317 |
fun schedule_of_provers provers num_slices = |
|
318 |
let |
|
75032
8d08bc7e8f98
tweak padding of prover slice schedule to include all provers
blanchet
parents:
75031
diff
changeset
|
319 |
val (known_provers, unknown_provers) = |
8d08bc7e8f98
tweak padding of prover slice schedule to include all provers
blanchet
parents:
75031
diff
changeset
|
320 |
List.partition (member (op =) default_slice_schedule) provers |
75029 | 321 |
|
75034 | 322 |
val default_slice_schedule = filter (member (op =) known_provers) default_slice_schedule |
323 |
val num_default_slices = length default_slice_schedule |
|
324 |
||
75029 | 325 |
fun round_robin _ [] = [] |
326 |
| round_robin 0 _ = [] |
|
327 |
| round_robin n (prover :: provers) = prover :: round_robin (n - 1) (provers @ [prover]) |
|
328 |
in |
|
329 |
if num_slices <= num_default_slices then |
|
330 |
take num_slices default_slice_schedule |
|
331 |
else |
|
75032
8d08bc7e8f98
tweak padding of prover slice schedule to include all provers
blanchet
parents:
75031
diff
changeset
|
332 |
default_slice_schedule |
8d08bc7e8f98
tweak padding of prover slice schedule to include all provers
blanchet
parents:
75031
diff
changeset
|
333 |
@ round_robin (num_slices - num_default_slices) (unknown_provers @ known_provers) |
75029 | 334 |
end |
335 |
||
75060
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
336 |
fun prover_slices_of_schedule ctxt factss |
75035 | 337 |
({max_facts, fact_filter, type_enc, lam_trans, uncurried_aliases, ...} : params) schedule = |
75029 | 338 |
let |
339 |
fun triplicate_slices original = |
|
340 |
let |
|
341 |
val shift = |
|
75339
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
342 |
map (apfst (fn (slice_size, num_facts, fact_filter) => |
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
343 |
(slice_size, num_facts, |
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
344 |
if fact_filter = mashN then mepoN |
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
345 |
else if fact_filter = mepoN then meshN |
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
346 |
else mashN))) |
75029 | 347 |
|
348 |
val shifted_once = shift original |
|
349 |
val shifted_twice = shift shifted_once |
|
350 |
in |
|
351 |
original @ shifted_once @ shifted_twice |
|
352 |
end |
|
353 |
||
75063
7ff39293e265
added possibility of extra options to SMT slices
blanchet
parents:
75060
diff
changeset
|
354 |
fun adjust_extra (ATP_Slice (format0, type_enc0, lam_trans0, uncurried_aliases0, |
7ff39293e265
added possibility of extra options to SMT slices
blanchet
parents:
75060
diff
changeset
|
355 |
extra_extra0)) = |
7ff39293e265
added possibility of extra options to SMT slices
blanchet
parents:
75060
diff
changeset
|
356 |
ATP_Slice (format0, the_default type_enc0 type_enc, the_default lam_trans0 lam_trans, |
7ff39293e265
added possibility of extra options to SMT slices
blanchet
parents:
75060
diff
changeset
|
357 |
the_default uncurried_aliases0 uncurried_aliases, extra_extra0) |
7ff39293e265
added possibility of extra options to SMT slices
blanchet
parents:
75060
diff
changeset
|
358 |
| adjust_extra (extra as SMT_Slice _) = extra |
75034 | 359 |
|
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
360 |
fun adjust_slice max_slice_size ((slice_size0, num_facts0, fact_filter0), extra) = |
75060
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
361 |
let |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
362 |
val slice_size = Int.min (max_slice_size, slice_size0) |
75060
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
363 |
val fact_filter = fact_filter |> the_default fact_filter0 |
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
364 |
val max_facts = max_facts |> the_default num_facts0 |
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
365 |
val num_facts = Int.min (max_facts, length (facts_of_filter fact_filter factss)) |
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
366 |
in |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
367 |
((slice_size, num_facts, fact_filter), adjust_extra extra) |
75060
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
368 |
end |
75034 | 369 |
|
75029 | 370 |
val provers = distinct (op =) schedule |
371 |
val prover_slices = |
|
75034 | 372 |
map (fn prover => (prover, |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
373 |
(is_none fact_filter ? triplicate_slices) (get_slices ctxt prover))) |
75034 | 374 |
provers |
75029 | 375 |
|
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
376 |
val max_threads = Multithreading.max_threads () |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
377 |
|
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
378 |
fun translate_schedule _ 0 _ = [] |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
379 |
| translate_schedule _ _ [] = [] |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
380 |
| translate_schedule prover_slices slices_left (prover :: schedule) = |
75029 | 381 |
(case AList.lookup (op =) prover_slices prover of |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
382 |
SOME (slice0 :: slices) => |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
383 |
let |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
384 |
val prover_slices' = AList.update (op =) (prover, slices) prover_slices |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
385 |
val slice as ((slice_size, _, _), _) = |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
386 |
adjust_slice ((slices_left + max_threads - 1) div max_threads) slice0 |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
387 |
in |
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
388 |
(prover, slice) :: translate_schedule prover_slices' (slices_left - slice_size) schedule |
75029 | 389 |
end |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
390 |
| _ => translate_schedule prover_slices slices_left schedule) |
75029 | 391 |
in |
75340
e1aa703c8cce
second and last step in making time slicing more flexible in Sledgehammer: try to honor desired slice size
blanchet
parents:
75339
diff
changeset
|
392 |
translate_schedule prover_slices (length schedule) schedule |
75034 | 393 |
|> distinct (op =) |
75029 | 394 |
end |
395 |
||
75031 | 396 |
fun run_sledgehammer (params as {verbose, spy, provers, induction_rules, max_facts, max_proofs, |
397 |
slices, ...}) |
|
75030 | 398 |
mode writeln_result i (fact_override as {only, ...}) state = |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
399 |
if null provers then |
63692 | 400 |
error "No prover is set" |
55286 | 401 |
else |
402 |
(case subgoal_count state of |
|
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
403 |
0 => (error "No subgoal!"; (false, (SH_None, ""))) |
55286 | 404 |
| n => |
405 |
let |
|
406 |
val _ = Proof.assert_backward state |
|
62735 | 407 |
val print = if mode = Normal andalso is_none writeln_result then writeln else K () |
408 |
||
75031 | 409 |
val found_proofs = Synchronized.var "found_proofs" 0 |
75030 | 410 |
|
411 |
fun found_proof prover_name = |
|
412 |
if mode = Normal then |
|
75031 | 413 |
(Synchronized.change found_proofs (fn n => n + 1); |
75076 | 414 |
(the_default writeln writeln_result) (prover_name ^ " found a proof...")) |
75030 | 415 |
else |
416 |
() |
|
62735 | 417 |
|
55286 | 418 |
val ctxt = Proof.context_of state |
74950
b350a1f2115d
added nearly_all_facts_of_context and uniformized its usage in Sledgehammer and Mirabelle
desharna
parents:
73975
diff
changeset
|
419 |
val inst_inducts = induction_rules = SOME Instantiate |
b350a1f2115d
added nearly_all_facts_of_context and uniformized its usage in Sledgehammer and Mirabelle
desharna
parents:
73975
diff
changeset
|
420 |
val {facts = chained_thms, goal, ...} = Proof.goal state |
55286 | 421 |
val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt |
422 |
val _ = |
|
423 |
(case find_first (not o is_prover_supported ctxt) provers of |
|
63692 | 424 |
SOME name => error ("No such prover: " ^ name) |
55286 | 425 |
| NONE => ()) |
426 |
val _ = print "Sledgehammering..." |
|
57037 | 427 |
val _ = spying spy (fn () => (state, i, "***", "Starting " ^ str_of_mode mode ^ " mode")) |
75002 | 428 |
val ({elapsed, ...}, all_facts) = Timing.timing |
429 |
(nearly_all_facts_of_context ctxt inst_inducts fact_override chained_thms hyp_ts) concl_t |
|
430 |
val _ = spying spy (fn () => (state, i, "All", |
|
431 |
"Extracting " ^ string_of_int (length all_facts) ^ " facts from background theory in " ^ |
|
432 |
string_of_int (Time.toMilliseconds elapsed) ^ " ms")) |
|
54090
a28992e35032
run relevance filter only once for ATPs and SMT solvers, since it should now yield the same results anyway
blanchet
parents:
54063
diff
changeset
|
433 |
|
55286 | 434 |
val spying_str_of_factss = |
435 |
commas o map (fn (filter, facts) => filter ^ ": " ^ string_of_int (length facts)) |
|
53800 | 436 |
|
55286 | 437 |
fun get_factss provers = |
438 |
let |
|
439 |
val max_max_facts = |
|
440 |
(case max_facts of |
|
441 |
SOME n => n |
|
442 |
| NONE => |
|
75029 | 443 |
fold (fn prover => |
75339
d9bb81999d2c
first step in making time slicing more flexible in Sledgehammer: label slices with 'slice size'
blanchet
parents:
75076
diff
changeset
|
444 |
fold (fn ((_, n, _), _) => Integer.max n) (get_slices ctxt prover)) |
75029 | 445 |
provers 0) |
75034 | 446 |
* 51 div 50 (* some slack to account for filtering of induction facts below *) |
75027
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
447 |
|
74998
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
desharna
parents:
74953
diff
changeset
|
448 |
val ({elapsed, ...}, factss) = Timing.timing |
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
desharna
parents:
74953
diff
changeset
|
449 |
(relevant_facts ctxt params (hd provers) max_max_facts fact_override hyp_ts concl_t) |
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
desharna
parents:
74953
diff
changeset
|
450 |
all_facts |
75027
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
451 |
|
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
452 |
val induction_rules = the_default (if only then Include else Exclude) induction_rules |
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
453 |
val factss = map (apsnd (maybe_filter_out_induction_rules induction_rules)) factss |
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
454 |
|
74998
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
desharna
parents:
74953
diff
changeset
|
455 |
val () = spying spy (fn () => (state, i, "All", |
75002 | 456 |
"Filtering facts in " ^ string_of_int (Time.toMilliseconds elapsed) ^ |
457 |
" ms (MaSh algorithm: " ^ str_of_mash_algorithm (the_mash_algorithm ()) ^ ")")); |
|
74998
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
desharna
parents:
74953
diff
changeset
|
458 |
val () = if verbose then print (string_of_factss factss) else () |
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
desharna
parents:
74953
diff
changeset
|
459 |
val () = spying spy (fn () => |
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
desharna
parents:
74953
diff
changeset
|
460 |
(state, i, "All", "Selected facts: " ^ spying_str_of_factss factss)) |
55286 | 461 |
in |
74998
fe14ceff1cfd
added syping of fact filtering time to sledgehammer
desharna
parents:
74953
diff
changeset
|
462 |
factss |
55286 | 463 |
end |
53800 | 464 |
|
58892
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents:
58843
diff
changeset
|
465 |
fun launch_provers () = |
55286 | 466 |
let |
75060
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
467 |
val factss = get_factss provers |
55286 | 468 |
val problem = |
469 |
{comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n, |
|
75060
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
470 |
factss = factss, found_proof = found_proof} |
69706
6d6235b828fc
get rid of visibility in MaSh -- it slows it down more than it helps
blanchet
parents:
63697
diff
changeset
|
471 |
val learn = mash_learn_proof ctxt params (Thm.prop_of goal) |
75027
a8efa30c380d
take 'induction_rules' into consideration, as well as 'max_facts' even when 'only' is set
blanchet
parents:
75026
diff
changeset
|
472 |
val launch = launch_prover_and_preplay params mode writeln_result learn |
75029 | 473 |
|
474 |
val schedule = |
|
75036 | 475 |
if mode = Auto_Try then provers |
476 |
else schedule_of_provers provers slices |
|
75060
789e0e1a9e33
more precise slicing computation and output when not enough lemmas are available (e.g. with the 'only' syntax 'sledgehammer (lem1 lem2 lem3)')
blanchet
parents:
75056
diff
changeset
|
477 |
val prover_slices = prover_slices_of_schedule ctxt factss params schedule |
75037 | 478 |
|
479 |
val _ = |
|
480 |
if verbose then |
|
481 |
writeln ("Running " ^ commas (map fst prover_slices) ^ "...") |
|
482 |
else |
|
483 |
() |
|
55286 | 484 |
in |
485 |
if mode = Auto_Try then |
|
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
486 |
(SH_Unknown, "") |
75029 | 487 |
|> fold (fn (prover, slice) => |
75025 | 488 |
fn accum as (SH_Some _, _) => accum |
75029 | 489 |
| _ => launch problem slice prover) |
490 |
prover_slices |
|
55286 | 491 |
else |
74950
b350a1f2115d
added nearly_all_facts_of_context and uniformized its usage in Sledgehammer and Mirabelle
desharna
parents:
73975
diff
changeset
|
492 |
(learn chained_thms; |
75031 | 493 |
Par_List.map (fn (prover, slice) => |
494 |
if Synchronized.value found_proofs < max_proofs then |
|
495 |
launch problem slice prover |
|
496 |
else |
|
75056
04a4881ff0fd
propagate right result when enough proofs have been found
blanchet
parents:
75054
diff
changeset
|
497 |
(SH_None, "")) |
75031 | 498 |
prover_slices |
74953
aade20a03edb
tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents:
74952
diff
changeset
|
499 |
|> max_outcome) |
55286 | 500 |
end |
501 |
in |
|
75020
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
502 |
(launch_provers () |
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
503 |
handle Timeout.TIMEOUT _ => (SH_Timeout, "")) |
75054 | 504 |
|> `(fn (outcome, message) => |
75020
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
505 |
(case outcome of |
75076 | 506 |
SH_Some _ => (the_default writeln writeln_result "QED"; true) |
507 |
| SH_Unknown => (the_default writeln writeln_result message; false) |
|
508 |
| SH_Timeout => (the_default writeln writeln_result "No proof found"; false) |
|
509 |
| SH_None => (the_default writeln writeln_result |
|
510 |
(if message = "" then "No proof found" else "Error: " ^ message); |
|
75075 | 511 |
false))) |
75020
b087610592b4
rationalized output for forthcoming slicing model
blanchet
parents:
75019
diff
changeset
|
512 |
end) |
38044 | 513 |
|
28582 | 514 |
end; |