| author | blanchet | 
| Wed, 01 Mar 2023 08:00:51 +0100 | |
| changeset 77423 | 779faa014564 | 
| parent 77418 | a8458f0df4ee | 
| child 77428 | 7c76221baecb | 
| 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_prover_minimize.ML  | 
| 
31037
 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 
immler@in.tum.de 
parents:  
diff
changeset
 | 
2  | 
Author: Philipp Meyer, TU Muenchen  | 
| 
36370
 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 
blanchet 
parents: 
36369 
diff
changeset
 | 
3  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
31037
 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 
immler@in.tum.de 
parents:  
diff
changeset
 | 
4  | 
|
| 40977 | 5  | 
Minimization of fact list for Metis using external provers.  | 
| 
31037
 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 
immler@in.tum.de 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 
immler@in.tum.de 
parents:  
diff
changeset
 | 
7  | 
|
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
8  | 
signature SLEDGEHAMMER_PROVER_MINIMIZE =  | 
| 32525 | 9  | 
sig  | 
| 46340 | 10  | 
type stature = ATP_Problem_Generate.stature  | 
| 55287 | 11  | 
type proof_method = Sledgehammer_Proof_Methods.proof_method  | 
12  | 
type play_outcome = Sledgehammer_Proof_Methods.play_outcome  | 
|
| 55201 | 13  | 
type mode = Sledgehammer_Prover.mode  | 
14  | 
type params = Sledgehammer_Prover.params  | 
|
15  | 
type prover = Sledgehammer_Prover.prover  | 
|
| 75025 | 16  | 
type prover_slice = Sledgehammer_Prover.prover_slice  | 
| 35867 | 17  | 
|
| 55205 | 18  | 
val is_prover_supported : Proof.context -> string -> bool  | 
19  | 
val is_prover_installed : Proof.context -> string -> bool  | 
|
20  | 
val get_prover : Proof.context -> mode -> string -> prover  | 
|
| 75029 | 21  | 
val get_slices : Proof.context -> string -> prover_slice list  | 
| 55205 | 22  | 
|
| 
42646
 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 
blanchet 
parents: 
42579 
diff
changeset
 | 
23  | 
val binary_min_facts : int Config.T  | 
| 75431 | 24  | 
val minimize_facts : (thm list -> unit) -> string -> params -> prover_slice -> bool -> int ->  | 
25  | 
int -> Proof.state -> thm -> ((string * stature) * thm list) list ->  | 
|
| 57739 | 26  | 
((string * stature) * thm list) list option  | 
| 57750 | 27  | 
* ((unit -> (string * stature) list * (proof_method * play_outcome)) -> string)  | 
| 54503 | 28  | 
val get_minimizing_prover : Proof.context -> mode -> (thm list -> unit) -> string -> prover  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35865 
diff
changeset
 | 
29  | 
end;  | 
| 32525 | 30  | 
|
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
31  | 
structure Sledgehammer_Prover_Minimize : SLEDGEHAMMER_PROVER_MINIMIZE =  | 
| 
31037
 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 
immler@in.tum.de 
parents:  
diff
changeset
 | 
32  | 
struct  | 
| 
 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 
immler@in.tum.de 
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  | 
| 
39496
 
a52a4e4399c1
got caught once again by SML's pattern maching (ctor vs. var)
 
blanchet 
parents: 
39491 
diff
changeset
 | 
35  | 
open ATP_Proof  | 
| 46320 | 36  | 
open ATP_Problem_Generate  | 
| 55212 | 37  | 
open ATP_Proof_Reconstruct  | 
| 
36142
 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 
blanchet 
parents: 
36063 
diff
changeset
 | 
38  | 
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: 
48085 
diff
changeset
 | 
39  | 
open Sledgehammer_Fact  | 
| 55287 | 40  | 
open Sledgehammer_Proof_Methods  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
41  | 
open Sledgehammer_Isar  | 
| 72400 | 42  | 
open Sledgehammer_ATP_Systems  | 
| 55201 | 43  | 
open Sledgehammer_Prover  | 
| 55205 | 44  | 
open Sledgehammer_Prover_ATP  | 
| 58061 | 45  | 
open Sledgehammer_Prover_SMT  | 
| 55205 | 46  | 
|
47  | 
fun is_prover_supported ctxt =  | 
|
| 75036 | 48  | 
is_atp orf is_smt_prover ctxt  | 
| 55205 | 49  | 
|
| 75036 | 50  | 
fun is_prover_installed ctxt prover_name =  | 
51  | 
if is_atp prover_name then  | 
|
52  | 
let val thy = Proof_Context.theory_of ctxt in  | 
|
53  | 
is_atp_installed thy prover_name  | 
|
54  | 
end  | 
|
55  | 
else  | 
|
56  | 
is_smt_prover_installed ctxt prover_name  | 
|
| 55205 | 57  | 
|
58  | 
fun get_prover ctxt mode name =  | 
|
| 75036 | 59  | 
if is_atp name then run_atp mode name  | 
60  | 
else if is_smt_prover ctxt name then run_smt_solver mode name  | 
|
61  | 
  else error ("No such prover: " ^ name)
 | 
|
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35865 
diff
changeset
 | 
62  | 
|
| 75029 | 63  | 
fun get_slices ctxt name =  | 
| 75025 | 64  | 
let val thy = Proof_Context.theory_of ctxt in  | 
| 75036 | 65  | 
if is_atp name then  | 
| 
75063
 
7ff39293e265
added possibility of extra options to SMT slices
 
blanchet 
parents: 
75036 
diff
changeset
 | 
66  | 
map (apsnd ATP_Slice) (#good_slices (get_atp thy name ()) ctxt)  | 
| 75025 | 67  | 
else if is_smt_prover ctxt name then  | 
| 
75063
 
7ff39293e265
added possibility of extra options to SMT slices
 
blanchet 
parents: 
75036 
diff
changeset
 | 
68  | 
map (apsnd SMT_Slice) (SMT_Solver.good_slices ctxt name)  | 
| 75025 | 69  | 
else  | 
70  | 
      error ("No such prover: " ^ name)
 | 
|
71  | 
end  | 
|
72  | 
||
| 
36370
 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 
blanchet 
parents: 
36369 
diff
changeset
 | 
73  | 
(* wrapper for calling external prover *)  | 
| 31236 | 74  | 
|
| 
40061
 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 
blanchet 
parents: 
40060 
diff
changeset
 | 
75  | 
fun n_facts names =  | 
| 
38698
 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 
blanchet 
parents: 
38696 
diff
changeset
 | 
76  | 
let val n = length names in  | 
| 
40061
 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 
blanchet 
parents: 
40060 
diff
changeset
 | 
77  | 
string_of_int n ^ " fact" ^ plural_s n ^  | 
| 57781 | 78  | 
(if n > 0 then ": " ^ (names |> map fst |> sort string_ord |> space_implode " ") else "")  | 
| 
38092
 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 
blanchet 
parents: 
38084 
diff
changeset
 | 
79  | 
end  | 
| 
 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 
blanchet 
parents: 
38084 
diff
changeset
 | 
80  | 
|
| 58843 | 81  | 
fun print silent f = if silent then () else writeln (f ())  | 
| 
41091
 
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
 
blanchet 
parents: 
41090 
diff
changeset
 | 
82  | 
|
| 54824 | 83  | 
fun test_facts ({debug, verbose, overlord, spy, provers, max_mono_iters, max_new_mono_instances,
 | 
| 57245 | 84  | 
type_enc, strict, lam_trans, uncurried_aliases, isar_proofs, compress, try0, smt_proofs,  | 
| 
73939
 
9231ea46e041
promoted "sledgehammer_instantiate_inducts" to proper option "induction_rules"
 
desharna 
parents: 
72400 
diff
changeset
 | 
85  | 
minimize, preplay_timeout, induction_rules, ...} : params)  | 
| 77423 | 86  | 
(slice as ((_, _, check_inconsistent, _, fact_filter), slice_extra)) silent  | 
| 
77418
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77269 
diff
changeset
 | 
87  | 
(prover : prover) timeout i n state goal facts =  | 
| 31236 | 88  | 
let  | 
| 57054 | 89  | 
val _ = print silent (fn () => "Testing " ^ n_facts (map fst facts) ^  | 
90  | 
(if verbose then " (timeout: " ^ string_of_time timeout ^ ")" else "") ^ "...")  | 
|
91  | 
||
| 
51005
 
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
 
blanchet 
parents: 
50669 
diff
changeset
 | 
92  | 
val facts = facts |> maps (fn (n, ths) => map (pair n) ths)  | 
| 
38100
 
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
 
blanchet 
parents: 
38094 
diff
changeset
 | 
93  | 
val params =  | 
| 
61311
 
150aa3015c47
removed legacy asynchronous mode in Sledgehammer
 
blanchet 
parents: 
60924 
diff
changeset
 | 
94  | 
      {debug = debug, verbose = verbose, overlord = overlord, spy = spy, provers = provers,
 | 
| 77423 | 95  | 
type_enc = type_enc, abduce = SOME 0, check_inconsistent = SOME false, strict = strict,  | 
| 
77418
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77269 
diff
changeset
 | 
96  | 
lam_trans = lam_trans, uncurried_aliases = uncurried_aliases, learn = false,  | 
| 
77269
 
bc43f86c9598
added refute mode to Sledgehammer to find 'counterexamples'
 
blanchet 
parents: 
75664 
diff
changeset
 | 
97  | 
fact_filter = NONE, induction_rules = induction_rules, max_facts = SOME (length facts),  | 
| 
73939
 
9231ea46e041
promoted "sledgehammer_instantiate_inducts" to proper option "induction_rules"
 
desharna 
parents: 
72400 
diff
changeset
 | 
98  | 
fact_thresholds = (1.01, 1.01), max_mono_iters = max_mono_iters,  | 
| 75031 | 99  | 
max_new_mono_instances = max_new_mono_instances, max_proofs = 1,  | 
100  | 
isar_proofs = isar_proofs, compress = compress, try0 = try0, smt_proofs = smt_proofs,  | 
|
101  | 
minimize = minimize, slices = 1, timeout = timeout, preplay_timeout = preplay_timeout,  | 
|
102  | 
expect = ""}  | 
|
| 40065 | 103  | 
val problem =  | 
| 
54141
 
f57f8e7a879f
generate a comment storing the goal nickname in "learn_prover"
 
blanchet 
parents: 
54130 
diff
changeset
 | 
104  | 
      {comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
 | 
| 
77418
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77269 
diff
changeset
 | 
105  | 
       factss = [("", facts)], has_already_found_something = K false, found_something = K ()}
 | 
| 
58654
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
106  | 
    val result0 as {outcome = outcome0, used_facts, used_from, preferred_methss, run_time,
 | 
| 
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
107  | 
message} =  | 
| 
77418
 
a8458f0df4ee
implemented ad hoc abduction in Sledgehammer with E
 
blanchet 
parents: 
77269 
diff
changeset
 | 
108  | 
prover params problem ((1, false, false, length facts, fact_filter), slice_extra)  | 
| 
58654
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
109  | 
    val result as {outcome, ...} =
 | 
| 
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
110  | 
if is_none outcome0 andalso  | 
| 
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
111  | 
forall (member (fn ((s, _), ((s', _), _)) => s = s') used_from) used_facts then  | 
| 
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
112  | 
result0  | 
| 
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
113  | 
else  | 
| 
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
114  | 
        {outcome = SOME MaybeUnprovable, used_facts = [], used_from = used_from,
 | 
| 
 
3e1cad27fc2f
special treatment of extensionality in minimizer
 
blanchet 
parents: 
58494 
diff
changeset
 | 
115  | 
preferred_methss = preferred_methss, run_time = run_time, message = message}  | 
| 
36223
 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 
blanchet 
parents: 
36143 
diff
changeset
 | 
116  | 
in  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
117  | 
print silent (fn () =>  | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
118  | 
(case outcome of  | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
119  | 
SOME failure => string_of_atp_failure failure  | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
120  | 
| NONE =>  | 
| 77423 | 121  | 
"Found " ^ (if check_inconsistent then "inconsistency" else "proof") ^  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
122  | 
(if length used_facts = length facts then "" else " with " ^ n_facts used_facts) ^  | 
| 63692 | 123  | 
         " (" ^ string_of_time run_time ^ ")"));
 | 
| 
38092
 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 
blanchet 
parents: 
38084 
diff
changeset
 | 
124  | 
result  | 
| 
36223
 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 
blanchet 
parents: 
36143 
diff
changeset
 | 
125  | 
end  | 
| 31236 | 126  | 
|
| 45372 | 127  | 
(* Give the external prover some slack. The ATP gets further slack because the  | 
128  | 
Sledgehammer preprocessing time is included in the estimate below but isn't  | 
|
129  | 
part of the timeout. *)  | 
|
130  | 
val slack_msecs = 200  | 
|
131  | 
||
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
132  | 
fun new_timeout timeout run_time =  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
133  | 
Int.min (Time.toMilliseconds timeout, Time.toMilliseconds run_time + slack_msecs)  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
134  | 
|> Time.fromMilliseconds  | 
| 45372 | 135  | 
|
| 45367 | 136  | 
(* The linear algorithm usually outperforms the binary algorithm when over 60%  | 
137  | 
of the facts are actually needed. The binary algorithm is much more  | 
|
138  | 
appropriate for provers that cannot return the list of used facts and hence  | 
|
139  | 
returns all facts as used. Since we cannot know in advance how many facts are  | 
|
140  | 
actually needed, we heuristically set the threshold to 10 facts. *)  | 
|
| 
42646
 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 
blanchet 
parents: 
42579 
diff
changeset
 | 
141  | 
val binary_min_facts =  | 
| 69593 | 142  | 
Attrib.setup_config_int \<^binding>\<open>sledgehammer_minimize_binary_min_facts\<close> (K 20)  | 
| 40977 | 143  | 
|
| 45368 | 144  | 
fun linear_minimize test timeout result xs =  | 
145  | 
let  | 
|
146  | 
fun min _ [] p = p  | 
|
147  | 
| min timeout (x :: xs) (seen, result) =  | 
|
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
148  | 
(case test timeout (xs @ seen) of  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
149  | 
          result as {outcome = NONE, used_facts, run_time, ...} : prover_result =>
 | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
150  | 
min (new_timeout timeout run_time) (filter_used_facts true used_facts xs)  | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
151  | 
(filter_used_facts false used_facts seen, result)  | 
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
152  | 
| _ => min timeout xs (x :: seen, result))  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
153  | 
in  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
154  | 
min timeout xs ([], result)  | 
| 
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
155  | 
end  | 
| 38015 | 156  | 
|
| 45368 | 157  | 
fun binary_minimize test timeout result xs =  | 
| 40977 | 158  | 
let  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
159  | 
    fun min depth (result as {run_time, ...} : prover_result) sup (xs as _ :: _ :: _) =
 | 
| 41743 | 160  | 
let  | 
| 45367 | 161  | 
val (l0, r0) = chop (length xs div 2) xs  | 
| 41743 | 162  | 
(*  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
163  | 
          val _ = warning (replicate_string depth " " ^ "{ " ^ "sup: " ^ n_facts (map fst sup))
 | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
164  | 
val _ = warning (replicate_string depth " " ^ " " ^ "xs: " ^ n_facts (map fst xs))  | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
165  | 
val _ = warning (replicate_string depth " " ^ " " ^ "l0: " ^ n_facts (map fst l0))  | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
166  | 
val _ = warning (replicate_string depth " " ^ " " ^ "r0: " ^ n_facts (map fst r0))  | 
| 41743 | 167  | 
*)  | 
| 45367 | 168  | 
val depth = depth + 1  | 
| 45372 | 169  | 
val timeout = new_timeout timeout run_time  | 
| 41743 | 170  | 
in  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
171  | 
(case test timeout (sup @ l0) of  | 
| 45372 | 172  | 
            result as {outcome = NONE, used_facts, ...} =>
 | 
| 48798 | 173  | 
min depth result (filter_used_facts true used_facts sup)  | 
| 57781 | 174  | 
(filter_used_facts true used_facts l0)  | 
| 45367 | 175  | 
| _ =>  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
176  | 
(case test timeout (sup @ r0) of  | 
| 45367 | 177  | 
              result as {outcome = NONE, used_facts, ...} =>
 | 
| 48798 | 178  | 
min depth result (filter_used_facts true used_facts sup)  | 
| 57054 | 179  | 
(filter_used_facts true used_facts r0)  | 
| 45367 | 180  | 
| _ =>  | 
181  | 
let  | 
|
182  | 
val (sup_r0, (l, result)) = min depth result (sup @ r0) l0  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58843 
diff
changeset
 | 
183  | 
val (sup, r0) = (sup, r0) |> apply2 (filter_used_facts true (map fst sup_r0))  | 
| 45367 | 184  | 
val (sup_l, (r, result)) = min depth result (sup @ l) r0  | 
| 48798 | 185  | 
val sup = sup |> filter_used_facts true (map fst sup_l)  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
186  | 
in (sup, (l @ r, result)) end))  | 
| 40977 | 187  | 
end  | 
| 41743 | 188  | 
(*  | 
189  | 
|> tap (fn _ => warning (replicate_string depth " " ^ "}"))  | 
|
190  | 
*)  | 
|
| 45367 | 191  | 
| min _ result sup xs = (sup, (xs, result))  | 
192  | 
in  | 
|
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
193  | 
(case snd (min 0 result [] xs) of  | 
| 45372 | 194  | 
      ([x], result as {run_time, ...}) =>
 | 
195  | 
(case test (new_timeout timeout run_time) [] of  | 
|
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
196  | 
        result as {outcome = NONE, ...} => ([], result)
 | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
197  | 
| _ => ([x], result))  | 
| 
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
198  | 
| p => p)  | 
| 45367 | 199  | 
end  | 
| 40977 | 200  | 
|
| 75431 | 201  | 
fun minimize_facts do_learn prover_name (params as {learn, timeout, ...}) slice silent i n state
 | 
202  | 
goal facts =  | 
|
| 31236 | 203  | 
let  | 
| 
40941
 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 
blanchet 
parents: 
40553 
diff
changeset
 | 
204  | 
val ctxt = Proof.context_of state  | 
| 58085 | 205  | 
val prover = get_prover ctxt Minimize prover_name  | 
| 
48796
 
0f94b8b69e79
consider removing chained facts last, so that they're more likely to be kept
 
blanchet 
parents: 
48399 
diff
changeset
 | 
206  | 
val (chained, non_chained) = List.partition is_fact_chained facts  | 
| 
59355
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
207  | 
|
| 
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
208  | 
fun test timeout non_chained =  | 
| 75025 | 209  | 
test_facts params slice silent prover timeout i n state goal (chained @ non_chained)  | 
| 31236 | 210  | 
in  | 
| 
75020
 
b087610592b4
rationalized output for forthcoming slicing model
 
blanchet 
parents: 
75017 
diff
changeset
 | 
211  | 
(print silent (fn () => "Sledgehammer minimizer: " ^ prover_name);  | 
| 
59355
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
212  | 
(case test timeout non_chained of  | 
| 45371 | 213  | 
       result as {outcome = NONE, used_facts, run_time, ...} =>
 | 
| 38015 | 214  | 
let  | 
| 
59355
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
215  | 
val non_chained = filter_used_facts true used_facts non_chained  | 
| 45372 | 216  | 
val min =  | 
| 
59355
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
217  | 
if length non_chained >= Config.get ctxt binary_min_facts then binary_minimize  | 
| 54815 | 218  | 
else linear_minimize  | 
| 
59355
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
219  | 
         val (min_facts, {message, ...}) =
 | 
| 
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
220  | 
min test (new_timeout timeout run_time) result non_chained  | 
| 
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
221  | 
val min_facts_and_chained = chained @ min_facts  | 
| 48321 | 222  | 
in  | 
223  | 
print silent (fn () => cat_lines  | 
|
| 57750 | 224  | 
["Minimized to " ^ n_facts (map fst min_facts)] ^  | 
| 
59355
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
225  | 
(case length chained of  | 
| 57750 | 226  | 
0 => ""  | 
| 63692 | 227  | 
| n => " (plus " ^ string_of_int n ^ " chained)"));  | 
| 
59355
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
228  | 
(if learn then do_learn (maps snd min_facts_and_chained) else ());  | 
| 
 
533f6cfc04c0
don't minimize chained facts -- this leads to subtle failures, e.g. if a method succeeds without a chained fact but fails with it
 
blanchet 
parents: 
59058 
diff
changeset
 | 
229  | 
(SOME min_facts_and_chained, message)  | 
| 48321 | 230  | 
end  | 
| 
57734
 
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
 
blanchet 
parents: 
57732 
diff
changeset
 | 
231  | 
     | {outcome = SOME TimedOut, ...} =>
 | 
| 57738 | 232  | 
(NONE, fn _ =>  | 
| 
54816
 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 
blanchet 
parents: 
54815 
diff
changeset
 | 
233  | 
"Timeout: You can increase the time limit using the \"timeout\" option (e.g., \  | 
| 63692 | 234  | 
\timeout = " ^ string_of_int (10 + Time.toMilliseconds timeout div 1000) ^ "\")")  | 
| 57738 | 235  | 
     | {message, ...} => (NONE, (prefix "Prover error: " o message))))
 | 
| 75664 | 236  | 
handle ERROR msg => (NONE, fn _ => "Warning: " ^ msg)  | 
| 31236 | 237  | 
end  | 
238  | 
||
| 57732 | 239  | 
fun maybe_minimize mode do_learn name (params as {verbose, minimize, ...})
 | 
| 75431 | 240  | 
    ({state, goal, subgoal, subgoal_count, ...} : prover_problem) slice
 | 
| 57738 | 241  | 
    (result as {outcome, used_facts, used_from, preferred_methss, run_time, message}
 | 
| 
57734
 
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
 
blanchet 
parents: 
57732 
diff
changeset
 | 
242  | 
: prover_result) =  | 
| 
58494
 
ed380b9594b5
always minimize, to reinvoke the prover with nicer options and yield a nicer Isar proof (potentially -- cf. 'full_proof')
 
blanchet 
parents: 
58085 
diff
changeset
 | 
243  | 
if is_some outcome then  | 
| 
48250
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
244  | 
result  | 
| 
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
245  | 
else  | 
| 
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
246  | 
let  | 
| 57738 | 247  | 
val (used_facts, message) =  | 
| 
48250
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
248  | 
if minimize then  | 
| 75431 | 249  | 
minimize_facts do_learn name params slice  | 
| 54824 | 250  | 
(not verbose orelse (mode <> Normal andalso mode <> MaSh)) subgoal subgoal_count state  | 
| 
57734
 
18bb3e1ff6f6
rationalized preplaying by eliminating (now superfluous) laziness
 
blanchet 
parents: 
57732 
diff
changeset
 | 
251  | 
goal (filter_used_facts true used_facts (map (apsnd single) used_from))  | 
| 
48250
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
252  | 
|>> Option.map (map fst)  | 
| 
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
253  | 
else  | 
| 57738 | 254  | 
(SOME used_facts, message)  | 
| 
48250
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
255  | 
in  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
256  | 
(case used_facts of  | 
| 
48250
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
257  | 
SOME used_facts =>  | 
| 
60924
 
610794dff23c
tuned signature, in accordance to sortBy in Scala;
 
wenzelm 
parents: 
59355 
diff
changeset
 | 
258  | 
        {outcome = NONE, used_facts = sort_by fst used_facts, used_from = used_from,
 | 
| 57738 | 259  | 
preferred_methss = preferred_methss, run_time = run_time, message = message}  | 
| 
55202
 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 
blanchet 
parents: 
55201 
diff
changeset
 | 
260  | 
| NONE => result)  | 
| 
48250
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
261  | 
end  | 
| 
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
262  | 
|
| 75025 | 263  | 
fun get_minimizing_prover ctxt mode do_learn name params problem slice =  | 
264  | 
get_prover ctxt mode name params problem slice  | 
|
| 75431 | 265  | 
|> maybe_minimize mode do_learn name params problem slice  | 
| 
48250
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
48085 
diff
changeset
 | 
266  | 
|
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35865 
diff
changeset
 | 
267  | 
end;  |