| author | blanchet |
| Wed, 06 Jan 2016 13:04:31 +0100 | |
| changeset 62081 | fd18b51bdc55 |
| parent 61330 | 20af2ad9261e |
| child 62219 | dbac573b27e7 |
| 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 |
|
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
18 |
|
|
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
19 |
val someN : string |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
20 |
val noneN : string |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
21 |
val timeoutN : string |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
22 |
val unknownN : string |
| 55212 | 23 |
|
|
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
|
24 |
val play_one_line_proof : bool -> Time.time -> (string * stature) list -> Proof.state -> int -> |
|
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
|
25 |
proof_method * proof_method list list -> (string * stature) list * (proof_method * play_outcome) |
| 51010 | 26 |
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
|
27 |
val run_sledgehammer : params -> mode -> (string -> unit) option -> int -> fact_override -> |
|
58892
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents:
58843
diff
changeset
|
28 |
Proof.state -> bool * (string * string list) |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
29 |
end; |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
30 |
|
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
31 |
structure Sledgehammer : SLEDGEHAMMER = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
32 |
struct |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
33 |
|
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43064
diff
changeset
|
34 |
open ATP_Util |
| 55212 | 35 |
open ATP_Proof |
| 46320 | 36 |
open ATP_Problem_Generate |
| 38023 | 37 |
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
|
38 |
open Sledgehammer_Fact |
| 55287 | 39 |
open Sledgehammer_Proof_Methods |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
40 |
open Sledgehammer_Isar_Proof |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
41 |
open Sledgehammer_Isar_Preplay |
| 57740 | 42 |
open Sledgehammer_Isar_Minimize |
| 55201 | 43 |
open Sledgehammer_Prover |
| 55212 | 44 |
open Sledgehammer_Prover_ATP |
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55201
diff
changeset
|
45 |
open Sledgehammer_Prover_Minimize |
| 48381 | 46 |
open Sledgehammer_MaSh |
|
40072
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents:
40071
diff
changeset
|
47 |
|
|
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
48 |
val someN = "some" |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
49 |
val noneN = "none" |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
50 |
val timeoutN = "timeout" |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
51 |
val unknownN = "unknown" |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
52 |
|
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
53 |
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
|
54 |
|
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
55 |
fun max_outcome_code codes = |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
56 |
NONE |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
57 |
|> fold (fn candidate => |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
58 |
fn accum as SOME _ => accum |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
59 |
| NONE => if member (op =) codes candidate then SOME candidate else NONE) |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
60 |
ordered_outcome_codes |
|
43020
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
61 |
|> the_default unknownN |
|
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents:
43006
diff
changeset
|
62 |
|
|
57774
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
63 |
fun is_metis_method (Metis_Method _) = true |
|
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
64 |
| is_metis_method _ = false |
|
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
65 |
|
| 60549 | 66 |
fun add_chained [] t = t |
67 |
| add_chained chained ((t as Const (@{const_name Pure.all}, _)) $ Abs (s, T, u)) =
|
|
68 |
t $ Abs (s, T, add_chained chained u) |
|
69 |
| add_chained chained t = Logic.list_implies (chained, t) |
|
70 |
||
|
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
|
71 |
fun play_one_line_proof minimize timeout used_facts0 state i (preferred_meth, methss) = |
|
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
|
72 |
let val used_facts = filter_out (fn (_, (sc, _)) => sc = Chained) used_facts0 in |
|
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
|
73 |
if timeout = Time.zeroTime then |
|
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
|
74 |
(used_facts, (preferred_meth, Play_Timed_Out Time.zeroTime)) |
|
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
|
75 |
else |
|
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
|
76 |
let |
|
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
|
77 |
val ctxt = Proof.context_of state |
|
60548
e6adb8868478
use right context for preplay, to avoid errors in fact lookup
blanchet
parents:
59582
diff
changeset
|
78 |
|
|
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
|
79 |
val fact_names = map fst used_facts |
|
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
|
80 |
val {facts = chained, goal, ...} = Proof.goal state
|
|
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
|
81 |
val goal_t = Logic.get_goal (Thm.prop_of goal) i |
|
61330
20af2ad9261e
further improved fine point w.r.t. replaying in the presence of chained facts and a non-empty meta-quantifier prefix + avoid printing internal names in backquotes
blanchet
parents:
61312
diff
changeset
|
82 |
|> add_chained (map (Thm.prop_of o forall_intr_vars) chained) |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
83 |
|
|
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
|
84 |
fun try_methss [] [] = (used_facts, (preferred_meth, Play_Timed_Out Time.zeroTime)) |
|
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
|
85 |
| try_methss ress [] = |
|
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
|
86 |
(used_facts, |
|
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
|
87 |
(case AList.lookup (op =) ress preferred_meth of |
|
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
|
88 |
SOME play => (preferred_meth, play) |
|
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
|
89 |
| NONE => hd (sort (play_outcome_ord o apply2 snd) (rev ress)))) |
|
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
|
90 |
| try_methss ress (meths :: methss) = |
|
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
|
91 |
let |
|
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
|
92 |
fun mk_step fact_names meths = |
|
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
|
93 |
Prove ([], [], ("", 0), goal_t, [], ([], fact_names), meths, "")
|
|
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
|
94 |
in |
|
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
|
95 |
(case preplay_isar_step ctxt timeout [] (mk_step fact_names meths) of |
|
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
|
96 |
(res as (meth, Played time)) :: _ => |
|
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
|
97 |
(* if a fact is needed by an ATP, it will be needed by "metis" *) |
|
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
|
98 |
if not minimize orelse is_metis_method meth then |
|
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
|
99 |
(used_facts, res) |
|
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
|
100 |
else |
|
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
|
101 |
let |
|
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
|
102 |
val (time', used_names') = |
|
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
|
103 |
minimized_isar_step ctxt time (mk_step fact_names [meth]) |
|
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
|
104 |
||> (facts_of_isar_step #> snd) |
|
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
|
105 |
val used_facts' = filter (member (op =) used_names' o fst) used_facts |
|
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
|
106 |
in |
|
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
|
107 |
(used_facts', (meth, Played time')) |
|
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
|
108 |
end |
|
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
|
109 |
| ress' => try_methss (ress' @ ress) methss) |
|
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
|
110 |
end |
|
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
|
111 |
in |
|
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
|
112 |
try_methss [] methss |
|
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
|
113 |
end |
|
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
|
114 |
end |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
115 |
|
|
61311
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
116 |
fun launch_prover (params as {debug, verbose, spy, max_facts, minimize, timeout, preplay_timeout,
|
|
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
117 |
expect, ...}) mode writeln_result only learn |
|
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54799
diff
changeset
|
118 |
{comment, state, goal, subgoal, subgoal_count, factss as (_, facts) :: _} name =
|
| 41089 | 119 |
let |
120 |
val ctxt = Proof.context_of state |
|
| 53800 | 121 |
|
|
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54799
diff
changeset
|
122 |
val hard_timeout = time_mult 3.0 timeout |
| 54062 | 123 |
val _ = spying spy (fn () => (state, subgoal, name, "Launched")); |
|
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
|
124 |
val max_facts = max_facts |> the_default (default_max_facts_of_prover ctxt name) |
| 48293 | 125 |
val num_facts = length facts |> not only ? Integer.min max_facts |
| 53800 | 126 |
|
| 41089 | 127 |
val problem = |
|
54141
f57f8e7a879f
generate a comment storing the goal nickname in "learn_prover"
blanchet
parents:
54126
diff
changeset
|
128 |
{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
|
129 |
subgoal_count = subgoal_count, |
| 53800 | 130 |
factss = factss |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
131 |
|> map (apsnd ((not (is_ho_atp ctxt name) |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
132 |
? filter_out (fn ((_, (_, Induction)), _) => true | _ => false)) |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
133 |
#> take num_facts))} |
| 53800 | 134 |
|
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51008
diff
changeset
|
135 |
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
|
136 |
tag_list 1 used_from |
|
51005
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents:
51004
diff
changeset
|
137 |
|> map (fn (j, fact) => fact |> apsnd (K j)) |
| 48798 | 138 |
|> 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
|
139 |
|> 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
|
140 |
|> 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
|
141 |
|> 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
|
142 |
" proof (of " ^ string_of_int (length facts) ^ "): ") "." |
| 58843 | 143 |
|> writeln |
| 53800 | 144 |
|
| 54062 | 145 |
fun spying_str_of_res ({outcome = NONE, used_facts, used_from, ...} : prover_result) =
|
146 |
let |
|
147 |
val num_used_facts = length used_facts |
|
| 54063 | 148 |
|
149 |
fun find_indices facts = |
|
150 |
tag_list 1 facts |
|
| 54062 | 151 |
|> map (fn (j, fact) => fact |> apsnd (K j)) |
152 |
|> filter_used_facts false used_facts |
|
| 54773 | 153 |
|> distinct (eq_fst (op =)) |
| 54062 | 154 |
|> map (prefix "@" o string_of_int o snd) |
| 54063 | 155 |
|
156 |
fun filter_info (fact_filter, facts) = |
|
157 |
let |
|
158 |
val indices = find_indices facts |
|
| 54773 | 159 |
(* "Int.max" is there for robustness -- it shouldn't be necessary *) |
160 |
val unknowns = replicate (Int.max (0, num_used_facts - length indices)) "?" |
|
161 |
in |
|
162 |
(commas (indices @ unknowns), fact_filter) |
|
163 |
end |
|
| 54063 | 164 |
|
165 |
val filter_infos = |
|
166 |
map filter_info (("actual", used_from) :: factss)
|
|
167 |
|> AList.group (op =) |
|
168 |
|> map (fn (indices, fact_filters) => commas fact_filters ^ ": " ^ indices) |
|
| 54062 | 169 |
in |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
170 |
"Success: Found proof with " ^ string_of_int num_used_facts ^ " of " ^ |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
171 |
string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ |
| 54063 | 172 |
(if num_used_facts = 0 then "" else ": " ^ commas filter_infos) |
| 53800 | 173 |
end |
174 |
| spying_str_of_res {outcome = SOME failure, ...} =
|
|
| 54062 | 175 |
"Failure: " ^ string_of_atp_failure failure |
| 53800 | 176 |
|
|
41255
a80024d7b71b
added debugging option to find out how good the relevance filter was at identifying relevant facts
blanchet
parents:
41245
diff
changeset
|
177 |
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
|
178 |
problem |
|
57735
056a55b44ec7
eliminated Sledgehammer's "min" subcommand (and lots of complications in the code)
blanchet
parents:
57734
diff
changeset
|
179 |
|> get_minimizing_prover ctxt mode learn name params |
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
180 |
|> verbose ? tap (fn {outcome = NONE, used_facts as _ :: _, used_from, ...} =>
|
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
181 |
print_used_facts used_facts used_from |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
182 |
| _ => ()) |
| 57056 | 183 |
|> spy ? tap (fn res => spying spy (fn () => (state, subgoal, name, spying_str_of_res res))) |
| 57739 | 184 |
|> (fn {outcome, used_facts, preferred_methss, message, ...} =>
|
|
57734
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
185 |
(if outcome = SOME ATP_Proof.TimedOut then timeoutN |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
186 |
else if is_some outcome then noneN |
|
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents:
57557
diff
changeset
|
187 |
else someN, |
|
57774
d2ad1320c770
honor 'dont_minimize' option when preplaying one-liner proof
blanchet
parents:
57755
diff
changeset
|
188 |
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
|
189 |
subgoal preferred_methss))) |
| 53800 | 190 |
|
| 41089 | 191 |
fun go () = |
192 |
let |
|
193 |
val (outcome_code, message) = |
|
194 |
if debug then |
|
195 |
really_go () |
|
196 |
else |
|
197 |
(really_go () |
|
| 57056 | 198 |
handle |
199 |
ERROR msg => (unknownN, fn () => "Error: " ^ msg ^ "\n") |
|
200 |
| exn => |
|
201 |
if Exn.is_interrupt exn then reraise exn |
|
202 |
else (unknownN, fn () => "Internal error:\n" ^ Runtime.exn_message exn ^ "\n")) |
|
203 |
||
| 41089 | 204 |
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
|
205 |
(* 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
|
206 |
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
|
207 |
machine. *) |
|
43e2b051339c
weaken the "expect" flag so that it doesn't trigger errors if a prover is not installed
blanchet
parents:
41138
diff
changeset
|
208 |
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
|
209 |
not (is_prover_installed ctxt name) then |
| 41089 | 210 |
() |
|
61311
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
211 |
else |
| 41089 | 212 |
error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
|
|
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
43004
diff
changeset
|
213 |
in (outcome_code, message) end |
| 41089 | 214 |
in |
| 43021 | 215 |
if mode = Auto_Try then |
|
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54799
diff
changeset
|
216 |
let val (outcome_code, message) = TimeLimit.timeLimit timeout go () in |
|
59184
830bb7ddb3ab
explicit message channels for "state", "information";
wenzelm
parents:
59058
diff
changeset
|
217 |
(outcome_code, if outcome_code = someN then [message ()] else []) |
| 41089 | 218 |
end |
|
61311
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
219 |
else |
| 43006 | 220 |
let |
221 |
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
|
222 |
val outcome = |
| 57056 | 223 |
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
|
224 |
val _ = |
|
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
61311
diff
changeset
|
225 |
if outcome <> "" andalso is_some writeln_result then the writeln_result outcome |
|
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
61311
diff
changeset
|
226 |
else writeln outcome |
|
58892
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents:
58843
diff
changeset
|
227 |
in (outcome_code, []) end |
| 41089 | 228 |
end |
229 |
||
| 48293 | 230 |
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
|
231 |
|
| 51008 | 232 |
fun string_of_facts facts = |
| 57384 | 233 |
"Including " ^ string_of_int (length facts) ^ " relevant fact" ^ plural_s (length facts) ^ ": " ^ |
| 51008 | 234 |
(facts |> map (fst o fst) |> space_implode " ") ^ "." |
235 |
||
| 51010 | 236 |
fun string_of_factss factss = |
| 57464 | 237 |
if forall (null o snd) factss then |
238 |
"Found no relevant facts." |
|
239 |
else |
|
240 |
cat_lines (map (fn (filter, facts) => |
|
241 |
(if filter = "" then "" else quote filter ^ ": ") ^ string_of_facts facts) factss) |
|
| 51008 | 242 |
|
|
61311
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
243 |
fun run_sledgehammer (params as {verbose, spy, provers, max_facts, ...}) mode writeln_result i
|
|
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
244 |
(fact_override as {only, ...}) state =
|
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
245 |
if null provers then |
|
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
246 |
error "No prover is set." |
| 55286 | 247 |
else |
248 |
(case subgoal_count state of |
|
|
61311
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
249 |
0 => (error "No subgoal!"; (false, (noneN, []))) |
| 55286 | 250 |
| n => |
251 |
let |
|
252 |
val _ = Proof.assert_backward state |
|
253 |
val print = |
|
| 61223 | 254 |
if mode = Normal andalso is_none writeln_result then writeln else K () |
| 55286 | 255 |
val ctxt = Proof.context_of state |
|
58928
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents:
58919
diff
changeset
|
256 |
val keywords = Thy_Header.get_keywords' ctxt |
| 55286 | 257 |
val {facts = chained, goal, ...} = Proof.goal state
|
258 |
val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt |
|
259 |
val ho_atp = exists (is_ho_atp ctxt) provers |
|
260 |
val css = clasimpset_rule_table_of ctxt |
|
261 |
val all_facts = |
|
| 58919 | 262 |
nearly_all_facts ctxt ho_atp fact_override keywords css chained hyp_ts concl_t |
| 55286 | 263 |
val _ = |
264 |
(case find_first (not o is_prover_supported ctxt) provers of |
|
265 |
SOME name => error ("No such prover: " ^ name ^ ".")
|
|
266 |
| NONE => ()) |
|
267 |
val _ = print "Sledgehammering..." |
|
| 57037 | 268 |
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
|
269 |
|
| 55286 | 270 |
val spying_str_of_factss = |
271 |
commas o map (fn (filter, facts) => filter ^ ": " ^ string_of_int (length facts)) |
|
| 53800 | 272 |
|
| 55286 | 273 |
fun get_factss provers = |
274 |
let |
|
275 |
val max_max_facts = |
|
276 |
(case max_facts of |
|
277 |
SOME n => n |
|
278 |
| NONE => |
|
279 |
0 |> fold (Integer.max o default_max_facts_of_prover ctxt) provers |
|
280 |
|> mode = Auto_Try ? (fn n => n div auto_try_max_facts_divisor)) |
|
281 |
val _ = spying spy (fn () => (state, i, "All", |
|
| 57557 | 282 |
"Filtering " ^ string_of_int (length all_facts) ^ " facts (MaSh algorithm: " ^ |
283 |
str_of_mash_algorithm (the_mash_algorithm ()) ^ ")")); |
|
| 55286 | 284 |
in |
285 |
all_facts |
|
286 |
|> relevant_facts ctxt params (hd provers) max_max_facts fact_override hyp_ts concl_t |
|
287 |
|> tap (fn factss => if verbose then print (string_of_factss factss) else ()) |
|
288 |
|> spy ? tap (fn factss => spying spy (fn () => |
|
289 |
(state, i, "All", "Selected facts: " ^ spying_str_of_factss factss))) |
|
290 |
end |
|
| 53800 | 291 |
|
|
58892
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents:
58843
diff
changeset
|
292 |
fun launch_provers () = |
| 55286 | 293 |
let |
294 |
val factss = get_factss provers |
|
295 |
val problem = |
|
296 |
{comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
|
|
297 |
factss = factss} |
|
| 59582 | 298 |
val learn = mash_learn_proof ctxt params (Thm.prop_of goal) all_facts |
| 61223 | 299 |
val launch = launch_prover params mode writeln_result only learn |
| 55286 | 300 |
in |
301 |
if mode = Auto_Try then |
|
|
58892
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents:
58843
diff
changeset
|
302 |
(unknownN, []) |
| 55286 | 303 |
|> fold (fn prover => fn accum as (outcome_code, _) => |
| 57056 | 304 |
if outcome_code = someN then accum else launch problem prover) |
305 |
provers |
|
| 55286 | 306 |
else |
|
57387
2b6fe2a48352
reintroduced MaSh hints, this time as persistent creatures
blanchet
parents:
57384
diff
changeset
|
307 |
(learn chained; |
|
2b6fe2a48352
reintroduced MaSh hints, this time as persistent creatures
blanchet
parents:
57384
diff
changeset
|
308 |
provers |
|
61311
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
309 |
|> Par_List.map (launch problem #> fst) |
|
58892
20aa19ecf2cc
eliminated obsolete Proof.goal_message -- print outcome more directly;
wenzelm
parents:
58843
diff
changeset
|
310 |
|> max_outcome_code |> rpair []) |
| 55286 | 311 |
end |
312 |
in |
|
|
61311
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
313 |
launch_provers () |
|
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
314 |
handle TimeLimit.TimeOut => |
|
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
blanchet
parents:
61223
diff
changeset
|
315 |
(print "Sledgehammer ran out of time."; (unknownN, [])) |
| 55286 | 316 |
end |
317 |
|> `(fn (outcome_code, _) => outcome_code = someN)) |
|
| 38044 | 318 |
|
| 28582 | 319 |
end; |