author | blanchet |
Mon, 04 Aug 2014 12:52:48 +0200 | |
changeset 57777 | 563df8185d98 |
parent 57776 | 1111a9a328fe |
child 57778 | cf4215911f85 |
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 |
51008 | 11 |
type fact = Sledgehammer_Fact.fact |
48292 | 12 |
type fact_override = Sledgehammer_Fact.fact_override |
57755 | 13 |
type proof_method = Sledgehammer_Proof_Methods.proof_method |
14 |
type play_outcome = Sledgehammer_Proof_Methods.play_outcome |
|
55201 | 15 |
type mode = Sledgehammer_Prover.mode |
16 |
type params = Sledgehammer_Prover.params |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
17 |
|
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
18 |
val someN : string |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
19 |
val noneN : string |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
20 |
val timeoutN : string |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
21 |
val unknownN : string |
55212 | 22 |
|
57774
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
23 |
val play_one_line_proof : bool -> Time.time -> (string * 'a) list -> Proof.state -> int -> |
57755 | 24 |
proof_method * proof_method list list -> (string * 'a) list * (proof_method * play_outcome) |
51010 | 25 |
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
|
26 |
val run_sledgehammer : params -> mode -> (string -> unit) option -> int -> fact_override -> |
57735
056a55b44ec7
eliminated Sledgehammer's "min" subcommand (and lots of complications in the code)
blanchet
parents:
57734
diff
changeset
|
27 |
Proof.state -> bool * (string * Proof.state) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
28 |
end; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
29 |
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
30 |
structure Sledgehammer : SLEDGEHAMMER = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
31 |
struct |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
32 |
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
33 |
open ATP_Util |
55212 | 34 |
open ATP_Proof |
46320 | 35 |
open ATP_Problem_Generate |
38023 | 36 |
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
|
37 |
open Sledgehammer_Fact |
55287 | 38 |
open Sledgehammer_Proof_Methods |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
39 |
open Sledgehammer_Isar_Proof |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
40 |
open Sledgehammer_Isar_Preplay |
57740 | 41 |
open Sledgehammer_Isar_Minimize |
55201 | 42 |
open Sledgehammer_Prover |
55212 | 43 |
open Sledgehammer_Prover_ATP |
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
44 |
open Sledgehammer_Prover_Minimize |
48381 | 45 |
open Sledgehammer_MaSh |
40072
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents:
40071
diff
changeset
|
46 |
|
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
47 |
val someN = "some" |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
48 |
val noneN = "none" |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
49 |
val timeoutN = "timeout" |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
50 |
val unknownN = "unknown" |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
51 |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
52 |
val ordered_outcome_codes = [someN, unknownN, timeoutN, noneN] |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
53 |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
54 |
fun max_outcome_code codes = |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
55 |
NONE |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
56 |
|> fold (fn candidate => |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
57 |
fn accum as SOME _ => accum |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
58 |
| NONE => if member (op =) codes candidate then SOME candidate else NONE) |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
59 |
ordered_outcome_codes |
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
60 |
|> the_default unknownN |
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
61 |
|
57053
46000c075d07
shorten Sledgehammer output, as suggested by Andrei Popescu
blanchet
parents:
57037
diff
changeset
|
62 |
fun prover_description verbose name num_facts = |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48299
diff
changeset
|
63 |
(quote name, |
57053
46000c075d07
shorten Sledgehammer output, as suggested by Andrei Popescu
blanchet
parents:
57037
diff
changeset
|
64 |
if verbose then " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts else "") |
41089 | 65 |
|
57774
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
66 |
fun is_metis_method (Metis_Method _) = true |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
67 |
| is_metis_method _ = false |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
68 |
|
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
69 |
fun play_one_line_proof minimize timeout used_facts state i |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
70 |
(preferred, methss as (meth :: _) :: _) = |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
71 |
let |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
72 |
fun dont_know () = |
57739 | 73 |
(used_facts, |
74 |
(if exists (fn meths => member (op =) meths preferred) methss then preferred else meth, |
|
75 |
Play_Timed_Out Time.zeroTime)) |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
76 |
in |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
77 |
if timeout = Time.zeroTime then |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
78 |
dont_know () |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
79 |
else |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
80 |
let |
57776
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
blanchet
parents:
57774
diff
changeset
|
81 |
val fact_names = map fst used_facts |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
82 |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
83 |
val {context = ctxt, facts = chained, goal} = Proof.goal state |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
84 |
val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
85 |
val goal_t = Logic.list_implies (map prop_of chained @ hyp_ts, concl_t) |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
86 |
|
57777 | 87 |
fun try_methss [] [] = dont_know () |
88 |
| try_methss ress [] = (used_facts, hd (sort (play_outcome_ord o pairself snd) ress)) |
|
89 |
| try_methss ress (meths :: methss) = |
|
57740 | 90 |
let |
91 |
fun mk_step fact_names meths = |
|
92 |
Prove ([], [], ("", 0), goal_t, [], ([], fact_names), meths, "") |
|
93 |
in |
|
94 |
(case preplay_isar_step ctxt timeout [] (mk_step fact_names meths) of |
|
57774
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
95 |
(res as (meth, Played time)) :: _ => |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
96 |
(* if a fact is needed by an ATP, it will be needed by "metis" *) |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
97 |
if not minimize orelse is_metis_method meth then |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
98 |
(used_facts, res) |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
99 |
else |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
100 |
let |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
101 |
val (time', used_names') = |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
102 |
minimized_isar_step ctxt time (mk_step fact_names [meth]) |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
103 |
||> (facts_of_isar_step #> snd) |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
104 |
val used_facts' = filter (member (op =) used_names' o fst) used_facts |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
105 |
in |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
106 |
(used_facts', (meth, Played time')) |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
107 |
end |
57777 | 108 |
| ress' => try_methss (ress' @ ress) methss) |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
109 |
end |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
110 |
in |
57777 | 111 |
try_methss [] methss |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
112 |
end |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
113 |
end |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
114 |
|
57774
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
115 |
fun launch_prover (params as {debug, verbose, spy, blocking, max_facts, minimize, timeout, |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
116 |
preplay_timeout, expect, ...}) mode output_result only learn |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54799
diff
changeset
|
117 |
{comment, state, goal, subgoal, subgoal_count, factss as (_, facts) :: _} name = |
41089 | 118 |
let |
119 |
val ctxt = Proof.context_of state |
|
53800 | 120 |
|
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54799
diff
changeset
|
121 |
val hard_timeout = time_mult 3.0 timeout |
54062 | 122 |
val _ = spying spy (fn () => (state, subgoal, name, "Launched")); |
41089 | 123 |
val birth_time = Time.now () |
42850
c8709be8a40f
distinguish between a soft timeout (30 s by defalt) and a hard timeout (60 s), to let minimization-based provers (such as CVC3, Yices, and occasionally the other provers) do their job
blanchet
parents:
42646
diff
changeset
|
124 |
val death_time = Time.+ (birth_time, hard_timeout) |
54126
6675cdc0d1ae
if slicing is disabled, pick the maximum number of facts, not the number of facts in the last slice
blanchet
parents:
54090
diff
changeset
|
125 |
val max_facts = max_facts |> the_default (default_max_facts_of_prover ctxt name) |
48293 | 126 |
val num_facts = length facts |> not only ? Integer.min max_facts |
53800 | 127 |
|
41089 | 128 |
val problem = |
54141
f57f8e7a879f
generate a comment storing the goal nickname in "learn_prover"
blanchet
parents:
54126
diff
changeset
|
129 |
{comment = comment, state = state, goal = goal, subgoal = subgoal, |
47904
67663c968d70
distinguish between instantiated and uninstantiated inductions -- the latter are OK for first-order provers
blanchet
parents:
47531
diff
changeset
|
130 |
subgoal_count = subgoal_count, |
53800 | 131 |
factss = factss |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
132 |
|> map (apsnd ((not (is_ho_atp ctxt name) |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
133 |
? filter_out (fn ((_, (_, Induction)), _) => true | _ => false)) |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
134 |
#> take num_facts))} |
53800 | 135 |
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51008
diff
changeset
|
136 |
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
|
137 |
tag_list 1 used_from |
51005
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents:
51004
diff
changeset
|
138 |
|> map (fn (j, fact) => fact |> apsnd (K j)) |
48798 | 139 |
|> 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
|
140 |
|> 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
|
141 |
|> commas |
82fc8c956cdc
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
blanchet
parents:
48384
diff
changeset
|
142 |
|> enclose ("Fact" ^ plural_s (length facts) ^ " in " ^ quote name ^ |
82fc8c956cdc
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
blanchet
parents:
48384
diff
changeset
|
143 |
" proof (of " ^ string_of_int (length facts) ^ "): ") "." |
82fc8c956cdc
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
blanchet
parents:
48384
diff
changeset
|
144 |
|> Output.urgent_message |
53800 | 145 |
|
54062 | 146 |
fun spying_str_of_res ({outcome = NONE, used_facts, used_from, ...} : prover_result) = |
147 |
let |
|
148 |
val num_used_facts = length used_facts |
|
54063 | 149 |
|
150 |
fun find_indices facts = |
|
151 |
tag_list 1 facts |
|
54062 | 152 |
|> map (fn (j, fact) => fact |> apsnd (K j)) |
153 |
|> filter_used_facts false used_facts |
|
54773 | 154 |
|> distinct (eq_fst (op =)) |
54062 | 155 |
|> map (prefix "@" o string_of_int o snd) |
54063 | 156 |
|
157 |
fun filter_info (fact_filter, facts) = |
|
158 |
let |
|
159 |
val indices = find_indices facts |
|
54773 | 160 |
(* "Int.max" is there for robustness -- it shouldn't be necessary *) |
161 |
val unknowns = replicate (Int.max (0, num_used_facts - length indices)) "?" |
|
162 |
in |
|
163 |
(commas (indices @ unknowns), fact_filter) |
|
164 |
end |
|
54063 | 165 |
|
166 |
val filter_infos = |
|
167 |
map filter_info (("actual", used_from) :: factss) |
|
168 |
|> AList.group (op =) |
|
169 |
|> map (fn (indices, fact_filters) => commas fact_filters ^ ": " ^ indices) |
|
54062 | 170 |
in |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
171 |
"Success: Found proof with " ^ string_of_int num_used_facts ^ " of " ^ |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
172 |
string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ |
54063 | 173 |
(if num_used_facts = 0 then "" else ": " ^ commas filter_infos) |
53800 | 174 |
end |
175 |
| spying_str_of_res {outcome = SOME failure, ...} = |
|
54062 | 176 |
"Failure: " ^ string_of_atp_failure failure |
53800 | 177 |
|
41255
a80024d7b71b
added debugging option to find out how good the relevance filter was at identifying relevant facts
blanchet
parents:
41245
diff
changeset
|
178 |
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
|
179 |
problem |
57735
056a55b44ec7
eliminated Sledgehammer's "min" subcommand (and lots of complications in the code)
blanchet
parents:
57734
diff
changeset
|
180 |
|> get_minimizing_prover ctxt mode learn name params |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
181 |
|> verbose ? tap (fn {outcome = NONE, used_facts as _ :: _, used_from, ...} => |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
182 |
print_used_facts used_facts used_from |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
183 |
| _ => ()) |
57056 | 184 |
|> spy ? tap (fn res => spying spy (fn () => (state, subgoal, name, spying_str_of_res res))) |
57739 | 185 |
|> (fn {outcome, used_facts, preferred_methss, message, ...} => |
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
186 |
(if outcome = SOME ATP_Proof.TimedOut then timeoutN |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
187 |
else if is_some outcome then noneN |
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
188 |
else someN, |
57774
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
189 |
fn () => message (fn () => play_one_line_proof minimize preplay_timeout used_facts state |
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
190 |
subgoal preferred_methss))) |
53800 | 191 |
|
41089 | 192 |
fun go () = |
193 |
let |
|
194 |
val (outcome_code, message) = |
|
195 |
if debug then |
|
196 |
really_go () |
|
197 |
else |
|
198 |
(really_go () |
|
57056 | 199 |
handle |
200 |
ERROR msg => (unknownN, fn () => "Error: " ^ msg ^ "\n") |
|
201 |
| exn => |
|
202 |
if Exn.is_interrupt exn then reraise exn |
|
203 |
else (unknownN, fn () => "Internal error:\n" ^ Runtime.exn_message exn ^ "\n")) |
|
204 |
||
41089 | 205 |
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
|
206 |
(* 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
|
207 |
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
|
208 |
machine. *) |
43e2b051339c
weaken the "expect" flag so that it doesn't trigger errors if a prover is not installed
blanchet
parents:
41138
diff
changeset
|
209 |
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
|
210 |
not (is_prover_installed ctxt name) then |
41089 | 211 |
() |
212 |
else if blocking then |
|
213 |
error ("Unexpected outcome: " ^ quote outcome_code ^ ".") |
|
214 |
else |
|
215 |
warning ("Unexpected outcome: " ^ quote outcome_code ^ "."); |
|
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43004
diff
changeset
|
216 |
in (outcome_code, message) end |
41089 | 217 |
in |
43021 | 218 |
if mode = Auto_Try then |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54799
diff
changeset
|
219 |
let val (outcome_code, message) = TimeLimit.timeLimit timeout go () in |
43006 | 220 |
(outcome_code, |
221 |
state |
|
222 |
|> outcome_code = someN |
|
223 |
? Proof.goal_message (fn () => |
|
52643
34c29356930e
more explicit Markup.information for messages produced by "auto" tools;
wenzelm
parents:
52555
diff
changeset
|
224 |
Pretty.mark Markup.information (Pretty.str (message ())))) |
41089 | 225 |
end |
226 |
else if blocking then |
|
43006 | 227 |
let |
228 |
val (outcome_code, message) = TimeLimit.timeLimit hard_timeout go () |
|
53048
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
wenzelm
parents:
52997
diff
changeset
|
229 |
val outcome = |
57056 | 230 |
if outcome_code = someN orelse mode = Normal then quote name ^ ": " ^ message () else "" |
53048
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
wenzelm
parents:
52997
diff
changeset
|
231 |
val _ = |
53052
a0db255af8c5
sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
wenzelm
parents:
53048
diff
changeset
|
232 |
if outcome <> "" andalso is_some output_result then |
53048
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
wenzelm
parents:
52997
diff
changeset
|
233 |
the output_result outcome |
52908
3461985dcbc3
dockable window for Sledgehammer, based on asynchronous/parallel query operation;
wenzelm
parents:
52643
diff
changeset
|
234 |
else |
53048
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
wenzelm
parents:
52997
diff
changeset
|
235 |
outcome |
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
wenzelm
parents:
52997
diff
changeset
|
236 |
|> Async_Manager.break_into_chunks |
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
wenzelm
parents:
52997
diff
changeset
|
237 |
|> List.app Output.urgent_message |
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
wenzelm
parents:
52997
diff
changeset
|
238 |
in (outcome_code, state) end |
41089 | 239 |
else |
57053
46000c075d07
shorten Sledgehammer output, as suggested by Andrei Popescu
blanchet
parents:
57037
diff
changeset
|
240 |
(Async_Manager.thread SledgehammerN birth_time death_time |
46000c075d07
shorten Sledgehammer output, as suggested by Andrei Popescu
blanchet
parents:
57037
diff
changeset
|
241 |
(prover_description verbose name num_facts) |
55212 | 242 |
((fn (outcome_code, message) => (verbose orelse outcome_code = someN, message ())) o go); |
43006 | 243 |
(unknownN, state)) |
41089 | 244 |
end |
245 |
||
48293 | 246 |
val auto_try_max_facts_divisor = 2 (* FUDGE *) |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
247 |
|
51008 | 248 |
fun string_of_facts facts = |
57384 | 249 |
"Including " ^ string_of_int (length facts) ^ " relevant fact" ^ plural_s (length facts) ^ ": " ^ |
51008 | 250 |
(facts |> map (fst o fst) |> space_implode " ") ^ "." |
251 |
||
51010 | 252 |
fun string_of_factss factss = |
57464 | 253 |
if forall (null o snd) factss then |
254 |
"Found no relevant facts." |
|
255 |
else |
|
256 |
cat_lines (map (fn (filter, facts) => |
|
257 |
(if filter = "" then "" else quote filter ^ ": ") ^ string_of_facts facts) factss) |
|
51008 | 258 |
|
57368 | 259 |
fun run_sledgehammer (params as {verbose, spy, blocking, provers, max_facts, ...}) mode |
57735
056a55b44ec7
eliminated Sledgehammer's "min" subcommand (and lots of complications in the code)
blanchet
parents:
57734
diff
changeset
|
260 |
output_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
|
261 |
if null provers then |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
262 |
error "No prover is set." |
55286 | 263 |
else |
264 |
(case subgoal_count state of |
|
265 |
0 => |
|
52908
3461985dcbc3
dockable window for Sledgehammer, based on asynchronous/parallel query operation;
wenzelm
parents:
52643
diff
changeset
|
266 |
((if blocking then error else Output.urgent_message) "No subgoal!"; (false, (noneN, state))) |
55286 | 267 |
| n => |
268 |
let |
|
269 |
val _ = Proof.assert_backward state |
|
270 |
val print = |
|
271 |
if mode = Normal andalso is_none output_result then Output.urgent_message else K () |
|
272 |
val ctxt = Proof.context_of state |
|
273 |
val {facts = chained, goal, ...} = Proof.goal state |
|
274 |
val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt |
|
275 |
val ho_atp = exists (is_ho_atp ctxt) provers |
|
276 |
val reserved = reserved_isar_keyword_table () |
|
277 |
val css = clasimpset_rule_table_of ctxt |
|
278 |
val all_facts = |
|
57262 | 279 |
nearly_all_facts ctxt ho_atp fact_override reserved css chained hyp_ts concl_t |
55286 | 280 |
val _ = () |> not blocking ? kill_provers |
281 |
val _ = |
|
282 |
(case find_first (not o is_prover_supported ctxt) provers of |
|
283 |
SOME name => error ("No such prover: " ^ name ^ ".") |
|
284 |
| NONE => ()) |
|
285 |
val _ = print "Sledgehammering..." |
|
57037 | 286 |
val _ = spying spy (fn () => (state, i, "***", "Starting " ^ str_of_mode mode ^ " mode")) |
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
|
287 |
|
55286 | 288 |
val spying_str_of_factss = |
289 |
commas o map (fn (filter, facts) => filter ^ ": " ^ string_of_int (length facts)) |
|
53800 | 290 |
|
55286 | 291 |
fun get_factss provers = |
292 |
let |
|
293 |
val max_max_facts = |
|
294 |
(case max_facts of |
|
295 |
SOME n => n |
|
296 |
| NONE => |
|
297 |
0 |> fold (Integer.max o default_max_facts_of_prover ctxt) provers |
|
298 |
|> mode = Auto_Try ? (fn n => n div auto_try_max_facts_divisor)) |
|
299 |
val _ = spying spy (fn () => (state, i, "All", |
|
57557 | 300 |
"Filtering " ^ string_of_int (length all_facts) ^ " facts (MaSh algorithm: " ^ |
301 |
str_of_mash_algorithm (the_mash_algorithm ()) ^ ")")); |
|
55286 | 302 |
in |
303 |
all_facts |
|
304 |
|> relevant_facts ctxt params (hd provers) max_max_facts fact_override hyp_ts concl_t |
|
305 |
|> tap (fn factss => if verbose then print (string_of_factss factss) else ()) |
|
306 |
|> spy ? tap (fn factss => spying spy (fn () => |
|
307 |
(state, i, "All", "Selected facts: " ^ spying_str_of_factss factss))) |
|
308 |
end |
|
53800 | 309 |
|
55286 | 310 |
fun launch_provers state = |
311 |
let |
|
312 |
val factss = get_factss provers |
|
313 |
val problem = |
|
314 |
{comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n, |
|
315 |
factss = factss} |
|
316 |
val learn = mash_learn_proof ctxt params (prop_of goal) all_facts |
|
57735
056a55b44ec7
eliminated Sledgehammer's "min" subcommand (and lots of complications in the code)
blanchet
parents:
57734
diff
changeset
|
317 |
val launch = launch_prover params mode output_result only learn |
55286 | 318 |
in |
319 |
if mode = Auto_Try then |
|
320 |
(unknownN, state) |
|
321 |
|> fold (fn prover => fn accum as (outcome_code, _) => |
|
57056 | 322 |
if outcome_code = someN then accum else launch problem prover) |
323 |
provers |
|
55286 | 324 |
else |
57387
2b6fe2a48352
reintroduced MaSh hints, this time as persistent creatures
blanchet
parents:
57384
diff
changeset
|
325 |
(learn chained; |
2b6fe2a48352
reintroduced MaSh hints, this time as persistent creatures
blanchet
parents:
57384
diff
changeset
|
326 |
provers |
2b6fe2a48352
reintroduced MaSh hints, this time as persistent creatures
blanchet
parents:
57384
diff
changeset
|
327 |
|> (if blocking then Par_List.map else map) (launch problem #> fst) |
2b6fe2a48352
reintroduced MaSh hints, this time as persistent creatures
blanchet
parents:
57384
diff
changeset
|
328 |
|> max_outcome_code |> rpair state) |
55286 | 329 |
end |
330 |
in |
|
331 |
(if blocking then launch_provers state |
|
332 |
else (Future.fork (tap (fn () => launch_provers state)); (unknownN, state))) |
|
333 |
handle TimeLimit.TimeOut => (print "Sledgehammer ran out of time."; (unknownN, state)) |
|
334 |
end |
|
335 |
|> `(fn (outcome_code, _) => outcome_code = someN)) |
|
38044 | 336 |
|
28582 | 337 |
end; |