| author | wenzelm |
| Thu, 11 Jul 2013 14:56:58 +0200 | |
| changeset 52597 | a8a81453833d |
| parent 52556 | c8357085217c |
| child 52632 | 23393c31c7fe |
| permissions | -rw-r--r-- |
|
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_provers.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 |
|
|
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
6 |
Generic prover abstraction for Sledgehammer. |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
7 |
*) |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
8 |
|
|
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
9 |
signature SLEDGEHAMMER_PROVERS = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
10 |
sig |
| 40181 | 11 |
type failure = ATP_Proof.failure |
| 46340 | 12 |
type stature = ATP_Problem_Generate.stature |
| 46320 | 13 |
type type_enc = ATP_Problem_Generate.type_enc |
|
51005
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents:
50927
diff
changeset
|
14 |
type fact = Sledgehammer_Fact.fact |
| 52555 | 15 |
type reconstructor = Sledgehammer_Reconstructor.reconstructor |
16 |
type play = Sledgehammer_Reconstructor.play |
|
17 |
type minimize_command = Sledgehammer_Reconstructor.minimize_command |
|
|
39493
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents:
39492
diff
changeset
|
18 |
|
|
48392
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
blanchet
parents:
48376
diff
changeset
|
19 |
datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize |
| 43021 | 20 |
|
| 35969 | 21 |
type params = |
| 48321 | 22 |
{debug : bool,
|
23 |
verbose : bool, |
|
24 |
overlord : bool, |
|
25 |
blocking : bool, |
|
26 |
provers : string list, |
|
27 |
type_enc : string option, |
|
28 |
strict : bool, |
|
29 |
lam_trans : string option, |
|
30 |
uncurried_aliases : bool option, |
|
31 |
learn : bool, |
|
32 |
fact_filter : string option, |
|
33 |
max_facts : int option, |
|
34 |
fact_thresholds : real * real, |
|
35 |
max_mono_iters : int option, |
|
36 |
max_new_mono_instances : int option, |
|
|
51190
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents:
51186
diff
changeset
|
37 |
isar_proofs : bool option, |
|
51130
76d68444cd59
renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents:
51024
diff
changeset
|
38 |
isar_compress : real, |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52555
diff
changeset
|
39 |
isar_compress_degree : int, |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52555
diff
changeset
|
40 |
merge_timeout_slack : real, |
| 48321 | 41 |
slice : bool, |
42 |
minimize : bool option, |
|
| 50557 | 43 |
timeout : Time.time option, |
44 |
preplay_timeout : Time.time option, |
|
| 51879 | 45 |
preplay_trace : bool, |
| 48321 | 46 |
expect : string} |
|
39493
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents:
39492
diff
changeset
|
47 |
|
|
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
48 |
type relevance_fudge = |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
49 |
{local_const_multiplier : real,
|
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
50 |
worse_irrel_freq : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
51 |
higher_order_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
52 |
abs_rel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
53 |
abs_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
54 |
skolem_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
55 |
theory_const_rel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
56 |
theory_const_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
57 |
chained_const_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
58 |
intro_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
59 |
elim_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
60 |
simp_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
61 |
local_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
62 |
assum_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
63 |
chained_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
64 |
max_imperfect : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
65 |
max_imperfect_exp : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
66 |
threshold_divisor : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
67 |
ridiculous_threshold : real} |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
68 |
|
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
69 |
type prover_problem = |
| 48321 | 70 |
{state : Proof.state,
|
71 |
goal : thm, |
|
72 |
subgoal : int, |
|
73 |
subgoal_count : int, |
|
| 51010 | 74 |
factss : (string * fact list) list} |
|
39493
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents:
39492
diff
changeset
|
75 |
|
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
76 |
type prover_result = |
| 48321 | 77 |
{outcome : failure option,
|
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
78 |
used_facts : (string * stature) list, |
|
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
79 |
used_from : fact list, |
| 48321 | 80 |
run_time : Time.time, |
|
50669
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
81 |
preplay : play Lazy.lazy, |
| 48321 | 82 |
message : play -> string, |
83 |
message_tail : string} |
|
|
39493
cb2208f2c07d
move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents:
39492
diff
changeset
|
84 |
|
| 43051 | 85 |
type prover = |
| 45520 | 86 |
params -> ((string * string list) list -> string -> minimize_command) |
87 |
-> prover_problem -> prover_result |
|
| 35867 | 88 |
|
|
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
89 |
val dest_dir : string Config.T |
|
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
90 |
val problem_prefix : string Config.T |
| 48143 | 91 |
val completish : bool Config.T |
|
44592
54906b0337ab
flip logic of boolean option so it's off by default
blanchet
parents:
44586
diff
changeset
|
92 |
val atp_full_names : bool Config.T |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
93 |
val smt_triggers : bool Config.T |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
94 |
val smt_weights : bool Config.T |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
95 |
val smt_weight_min_facts : int Config.T |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
96 |
val smt_min_weight : int Config.T |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
97 |
val smt_max_weight : int Config.T |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
98 |
val smt_max_weight_index : int Config.T |
|
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
99 |
val smt_weight_curve : (int -> int) Unsynchronized.ref |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
100 |
val smt_max_slices : int Config.T |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
101 |
val smt_slice_fact_frac : real Config.T |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
102 |
val smt_slice_time_frac : real Config.T |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
103 |
val smt_slice_min_secs : int Config.T |
|
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48314
diff
changeset
|
104 |
val SledgehammerN : string |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
105 |
val plain_metis : reconstructor |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
106 |
val select_smt_solver : string -> Proof.context -> Proof.context |
| 45520 | 107 |
val extract_reconstructor : |
|
45561
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
108 |
params -> reconstructor -> string * (string * string list) list |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
109 |
val is_reconstructor : string -> bool |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
110 |
val is_atp : theory -> string -> bool |
|
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
111 |
val is_smt_prover : Proof.context -> string -> bool |
|
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47946
diff
changeset
|
112 |
val is_ho_atp: Proof.context -> string -> bool |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
113 |
val is_unit_equational_atp : Proof.context -> string -> bool |
|
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
114 |
val is_prover_supported : Proof.context -> string -> bool |
|
40072
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents:
40071
diff
changeset
|
115 |
val is_prover_installed : Proof.context -> string -> bool |
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
116 |
val remotify_prover_if_supported_and_not_already_remote : |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
117 |
Proof.context -> string -> string option |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
118 |
val remotify_prover_if_not_installed : |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
119 |
Proof.context -> string -> string option |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
120 |
val default_max_facts_of_prover : Proof.context -> bool -> string -> int |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
121 |
val is_unit_equality : term -> bool |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
122 |
val is_appropriate_prop_of_prover : Proof.context -> string -> term -> bool |
|
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
123 |
val is_built_in_const_of_prover : |
|
41336
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
124 |
Proof.context -> string -> string * typ -> term list -> bool * term list |
|
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
125 |
val atp_relevance_fudge : relevance_fudge |
|
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
126 |
val smt_relevance_fudge : relevance_fudge |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
127 |
val relevance_fudge_of_prover : Proof.context -> string -> relevance_fudge |
|
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
128 |
val weight_smt_fact : |
| 46340 | 129 |
Proof.context -> int -> ((string * stature) * thm) * int |
130 |
-> (string * stature) * (int option * thm) |
|
|
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
131 |
val supported_provers : Proof.context -> unit |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
132 |
val kill_provers : unit -> unit |
|
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
133 |
val running_provers : unit -> unit |
|
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
134 |
val messages : int option -> unit |
| 48798 | 135 |
val is_fact_chained : (('a * stature) * 'b) -> bool
|
136 |
val filter_used_facts : |
|
137 |
bool -> (''a * stature) list -> ((''a * stature) * 'b) list ->
|
|
138 |
((''a * stature) * 'b) list
|
|
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
139 |
val isar_proof_reconstructor : Proof.context -> string -> string |
| 43021 | 140 |
val get_prover : Proof.context -> mode -> string -> prover |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
141 |
end; |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
142 |
|
|
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
41066
diff
changeset
|
143 |
structure Sledgehammer_Provers : SLEDGEHAMMER_PROVERS = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
144 |
struct |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
145 |
|
|
43085
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
blanchet
parents:
43063
diff
changeset
|
146 |
open ATP_Util |
| 38028 | 147 |
open ATP_Problem |
|
39491
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents:
39453
diff
changeset
|
148 |
open ATP_Proof |
| 38028 | 149 |
open ATP_Systems |
| 46320 | 150 |
open ATP_Problem_Generate |
151 |
open ATP_Proof_Reconstruct |
|
| 45521 | 152 |
open Metis_Tactic |
| 38023 | 153 |
open Sledgehammer_Util |
|
51005
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents:
50927
diff
changeset
|
154 |
open Sledgehammer_Fact |
| 52555 | 155 |
open Sledgehammer_Reconstructor |
156 |
open Sledgehammer_Print |
|
|
49881
d9d73ebf9274
added proof minimization code from Steffen Smolka
blanchet
parents:
48802
diff
changeset
|
157 |
open Sledgehammer_Reconstruct |
|
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
158 |
|
|
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
159 |
|
|
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
160 |
(** The Sledgehammer **) |
|
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
161 |
|
|
48392
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
blanchet
parents:
48376
diff
changeset
|
162 |
datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize |
| 43021 | 163 |
|
| 45376 | 164 |
(* Identifier that distinguishes Sledgehammer from other tools that could use |
|
38102
019a49759829
fix bug in the newly introduced "bound concealing" code
blanchet
parents:
38100
diff
changeset
|
165 |
"Async_Manager". *) |
|
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48314
diff
changeset
|
166 |
val SledgehammerN = "Sledgehammer" |
| 37585 | 167 |
|
| 45520 | 168 |
val reconstructor_names = [metisN, smtN] |
| 46365 | 169 |
val plain_metis = Metis (hd partial_type_encs, combsN) |
|
45561
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
170 |
val is_reconstructor = member (op =) reconstructor_names |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43226
diff
changeset
|
171 |
|
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
172 |
val is_atp = member (op =) o supported_atps |
|
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
173 |
|
|
43233
2749c357f865
use the proper prover name, e.g. metis_full_types, not metis (full_types), for minimizing
blanchet
parents:
43232
diff
changeset
|
174 |
val select_smt_solver = Context.proof_map o SMT_Config.select_solver |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
175 |
|
| 45376 | 176 |
fun is_smt_prover ctxt = member (op =) (SMT_Solver.available_solvers_of ctxt) |
| 40062 | 177 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
178 |
fun is_atp_of_format is_format ctxt name = |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
179 |
let val thy = Proof_Context.theory_of ctxt in |
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
180 |
case try (get_atp thy) name of |
|
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47531
diff
changeset
|
181 |
SOME config => |
|
48716
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48656
diff
changeset
|
182 |
exists (fn (_, ((_, format, _, _, _), _)) => is_format format) |
|
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47531
diff
changeset
|
183 |
(#best_slices (config ()) ctxt) |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
184 |
| NONE => false |
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
185 |
end |
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
186 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
187 |
val is_unit_equational_atp = is_atp_of_format (curry (op =) CNF_UEQ) |
|
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
188 |
val is_ho_atp = is_atp_of_format is_format_higher_order |
| 44597 | 189 |
|
| 45376 | 190 |
fun is_prover_supported ctxt = |
| 42361 | 191 |
let val thy = Proof_Context.theory_of ctxt in |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
192 |
is_reconstructor orf is_atp thy orf is_smt_prover ctxt |
|
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
193 |
end |
|
40072
27f2a45b0aab
more robust handling of "remote_" vs. non-"remote_" provers
blanchet
parents:
40071
diff
changeset
|
194 |
|
|
41432
3214c39777ab
differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents:
41336
diff
changeset
|
195 |
fun is_prover_installed ctxt = |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
196 |
is_reconstructor orf is_smt_prover ctxt orf |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
197 |
is_atp_installed (Proof_Context.theory_of ctxt) |
|
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
198 |
|
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
199 |
fun remotify_prover_if_supported_and_not_already_remote ctxt name = |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
200 |
if String.isPrefix remote_prefix name then |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
201 |
SOME name |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
202 |
else |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
203 |
let val remote_name = remote_prefix ^ name in |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
204 |
if is_prover_supported ctxt remote_name then SOME remote_name else NONE |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
205 |
end |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
206 |
|
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
207 |
fun remotify_prover_if_not_installed ctxt name = |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
208 |
if is_prover_supported ctxt name andalso is_prover_installed ctxt name then |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
209 |
SOME name |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
210 |
else |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
211 |
remotify_prover_if_supported_and_not_already_remote ctxt name |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
212 |
|
| 45706 | 213 |
fun get_slices slice slices = |
214 |
(0 upto length slices - 1) ~~ slices |> not slice ? (List.last #> single) |
|
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
215 |
|
| 48293 | 216 |
val reconstructor_default_max_facts = 20 |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
217 |
|
|
51186
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
218 |
fun slice_max_facts (_, (_, ( ((max_facts, _), _, _, _, _), _))) = max_facts |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
219 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
220 |
fun default_max_facts_of_prover ctxt slice name = |
| 42361 | 221 |
let val thy = Proof_Context.theory_of ctxt in |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
222 |
if is_reconstructor name then |
| 48293 | 223 |
reconstructor_default_max_facts |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
224 |
else if is_atp thy name then |
|
51186
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
225 |
fold (Integer.max o slice_max_facts) |
|
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47531
diff
changeset
|
226 |
(get_slices slice (#best_slices (get_atp thy name ()) ctxt)) 0 |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
227 |
else (* is_smt_prover ctxt name *) |
|
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
228 |
SMT_Solver.default_max_relevant ctxt name |
|
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
229 |
end |
|
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
230 |
|
|
42956
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
231 |
fun is_if (@{const_name If}, _) = true
|
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
232 |
| is_if _ = false |
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
233 |
|
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
234 |
(* Beware of "if and only if" (which is translated as such) and "If" (which is |
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
235 |
translated to conditional equations). *) |
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
236 |
fun is_good_unit_equality T t u = |
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
237 |
T <> @{typ bool} andalso not (exists (exists_Const is_if) [t, u])
|
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
238 |
|
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
239 |
fun is_unit_equality (@{const Trueprop} $ t) = is_unit_equality t
|
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
240 |
| is_unit_equality (Const (@{const_name all}, _) $ Abs (_, _, t)) =
|
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
241 |
is_unit_equality t |
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
242 |
| is_unit_equality (Const (@{const_name All}, _) $ Abs (_, _, t)) =
|
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
243 |
is_unit_equality t |
|
42956
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
244 |
| is_unit_equality (Const (@{const_name "=="}, Type (_, [T, _])) $ t $ u) =
|
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
245 |
is_good_unit_equality T t u |
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
246 |
| is_unit_equality (Const (@{const_name HOL.eq}, Type (_ , [T, _])) $ t $ u) =
|
|
9aeb0f6ad971
filter Waldmeister facts better -- and don't encode type classes as predicates, since it doesn't like conditional equations
blanchet
parents:
42952
diff
changeset
|
247 |
is_good_unit_equality T t u |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
248 |
| is_unit_equality _ = false |
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
249 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
250 |
fun is_appropriate_prop_of_prover ctxt name = |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
251 |
if is_unit_equational_atp ctxt name then is_unit_equality else K true |
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
252 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
253 |
fun is_built_in_const_of_prover ctxt name = |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
254 |
if is_smt_prover ctxt name then |
|
41336
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
255 |
let val ctxt = ctxt |> select_smt_solver name in |
|
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
256 |
fn x => fn ts => |
|
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
257 |
if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then |
|
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
258 |
(true, []) |
|
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
259 |
else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then |
|
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
260 |
(true, ts) |
|
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
261 |
else |
|
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
262 |
(false, ts) |
|
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
263 |
end |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
264 |
else |
|
41336
0ea5b9c7d233
proper handling of the arguments of SMT builtins -- for numerals, ignore the arguments (Pls, Bit0, Bit1, ..), for functions, consider them;
blanchet
parents:
41335
diff
changeset
|
265 |
fn (s, _) => fn ts => (member (op =) atp_irrelevant_consts s, ts) |
|
40071
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
266 |
|
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
267 |
(* FUDGE *) |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
268 |
val atp_relevance_fudge = |
| 42738 | 269 |
{local_const_multiplier = 1.5,
|
|
41159
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
270 |
worse_irrel_freq = 100.0, |
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
271 |
higher_order_irrel_weight = 1.05, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
272 |
abs_rel_weight = 0.5, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
273 |
abs_irrel_weight = 2.0, |
|
47934
08d7aff8c7e6
lower skolem penalty to ensure that some useful facts with existentials, e.g. congruence of "setsum", eventually get picked up
blanchet
parents:
47912
diff
changeset
|
274 |
skolem_irrel_weight = 0.05, |
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
275 |
theory_const_rel_weight = 0.5, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
276 |
theory_const_irrel_weight = 0.25, |
|
42735
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
blanchet
parents:
42730
diff
changeset
|
277 |
chained_const_irrel_weight = 0.25, |
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
278 |
intro_bonus = 0.15, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
279 |
elim_bonus = 0.15, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
280 |
simp_bonus = 0.15, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
281 |
local_bonus = 0.55, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
282 |
assum_bonus = 1.05, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
283 |
chained_bonus = 1.5, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
284 |
max_imperfect = 11.5, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
285 |
max_imperfect_exp = 1.0, |
|
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
286 |
threshold_divisor = 2.0, |
| 41093 | 287 |
ridiculous_threshold = 0.01} |
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
288 |
|
|
40071
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
289 |
(* FUDGE (FIXME) *) |
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
290 |
val smt_relevance_fudge = |
| 42738 | 291 |
{local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
|
|
41159
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
292 |
worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge, |
|
40071
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
293 |
higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
294 |
abs_rel_weight = #abs_rel_weight atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
295 |
abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
296 |
skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
297 |
theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
298 |
theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge, |
|
42735
1d375de437e9
fine-tuned the relevance filter, so that equations of the form "c = (%x. _)" and constants occurring in chained facts are not unduely penalized
blanchet
parents:
42730
diff
changeset
|
299 |
chained_const_irrel_weight = #chained_const_irrel_weight atp_relevance_fudge, |
|
40071
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
300 |
intro_bonus = #intro_bonus atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
301 |
elim_bonus = #elim_bonus atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
302 |
simp_bonus = #simp_bonus atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
303 |
local_bonus = #local_bonus atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
304 |
assum_bonus = #assum_bonus atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
305 |
chained_bonus = #chained_bonus atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
306 |
max_imperfect = #max_imperfect atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
307 |
max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
308 |
threshold_divisor = #threshold_divisor atp_relevance_fudge, |
|
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
309 |
ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge} |
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
310 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
311 |
fun relevance_fudge_of_prover ctxt name = |
|
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
312 |
if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge |
|
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
313 |
|
|
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
314 |
fun supported_provers ctxt = |
|
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
315 |
let |
| 42361 | 316 |
val thy = Proof_Context.theory_of ctxt |
|
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
317 |
val (remote_provers, local_provers) = |
| 45520 | 318 |
reconstructor_names @ |
|
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
319 |
sort_strings (supported_atps thy) @ |
|
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
320 |
sort_strings (SMT_Solver.available_solvers_of ctxt) |
|
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
321 |
|> List.partition (String.isPrefix remote_prefix) |
|
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
322 |
in |
|
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41723
diff
changeset
|
323 |
Output.urgent_message ("Supported provers: " ^
|
|
40205
277508b07418
if "debug" is on, print list of relevant facts (poweruser request);
blanchet
parents:
40204
diff
changeset
|
324 |
commas (local_provers @ remote_provers) ^ ".") |
|
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
325 |
end |
| 35969 | 326 |
|
|
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48314
diff
changeset
|
327 |
fun kill_provers () = Async_Manager.kill_threads SledgehammerN "prover" |
|
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48314
diff
changeset
|
328 |
fun running_provers () = Async_Manager.running_threads SledgehammerN "prover" |
|
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48314
diff
changeset
|
329 |
val messages = Async_Manager.thread_messages SledgehammerN "prover" |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
330 |
|
|
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
331 |
|
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
332 |
(** problems, results, ATPs, etc. **) |
| 35969 | 333 |
|
334 |
type params = |
|
| 48321 | 335 |
{debug : bool,
|
336 |
verbose : bool, |
|
337 |
overlord : bool, |
|
338 |
blocking : bool, |
|
339 |
provers : string list, |
|
340 |
type_enc : string option, |
|
341 |
strict : bool, |
|
342 |
lam_trans : string option, |
|
343 |
uncurried_aliases : bool option, |
|
344 |
learn : bool, |
|
345 |
fact_filter : string option, |
|
346 |
max_facts : int option, |
|
347 |
fact_thresholds : real * real, |
|
348 |
max_mono_iters : int option, |
|
349 |
max_new_mono_instances : int option, |
|
|
51190
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents:
51186
diff
changeset
|
350 |
isar_proofs : bool option, |
|
51130
76d68444cd59
renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents:
51024
diff
changeset
|
351 |
isar_compress : real, |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52555
diff
changeset
|
352 |
isar_compress_degree : int, |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52555
diff
changeset
|
353 |
merge_timeout_slack : real, |
| 48321 | 354 |
slice : bool, |
355 |
minimize : bool option, |
|
| 50557 | 356 |
timeout : Time.time option, |
357 |
preplay_timeout : Time.time option, |
|
| 51879 | 358 |
preplay_trace : bool, |
| 48321 | 359 |
expect : string} |
| 35867 | 360 |
|
|
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
361 |
type relevance_fudge = |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
362 |
{local_const_multiplier : real,
|
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
363 |
worse_irrel_freq : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
364 |
higher_order_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
365 |
abs_rel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
366 |
abs_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
367 |
skolem_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
368 |
theory_const_rel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
369 |
theory_const_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
370 |
chained_const_irrel_weight : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
371 |
intro_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
372 |
elim_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
373 |
simp_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
374 |
local_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
375 |
assum_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
376 |
chained_bonus : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
377 |
max_imperfect : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
378 |
max_imperfect_exp : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
379 |
threshold_divisor : real, |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
380 |
ridiculous_threshold : real} |
|
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48250
diff
changeset
|
381 |
|
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
382 |
type prover_problem = |
| 48321 | 383 |
{state : Proof.state,
|
384 |
goal : thm, |
|
385 |
subgoal : int, |
|
386 |
subgoal_count : int, |
|
| 51010 | 387 |
factss : (string * fact list) list} |
| 35867 | 388 |
|
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
389 |
type prover_result = |
| 48321 | 390 |
{outcome : failure option,
|
391 |
used_facts : (string * stature) list, |
|
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
392 |
used_from : fact list, |
| 48321 | 393 |
run_time : Time.time, |
|
50669
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
394 |
preplay : play Lazy.lazy, |
| 48321 | 395 |
message : play -> string, |
396 |
message_tail : string} |
|
| 35867 | 397 |
|
| 43051 | 398 |
type prover = |
| 45520 | 399 |
params -> ((string * string list) list -> string -> minimize_command) |
400 |
-> prover_problem -> prover_result |
|
| 35867 | 401 |
|
| 38023 | 402 |
(* configuration attributes *) |
403 |
||
|
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
404 |
(* Empty string means create files in Isabelle's temporary files directory. *) |
|
42616
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42593
diff
changeset
|
405 |
val dest_dir = |
|
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42593
diff
changeset
|
406 |
Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
|
|
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42593
diff
changeset
|
407 |
val problem_prefix = |
|
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents:
42593
diff
changeset
|
408 |
Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
|
| 48143 | 409 |
val completish = |
410 |
Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
|
|
| 28484 | 411 |
|
|
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
412 |
(* In addition to being easier to read, readable names are often much shorter, |
|
44394
20bd9f90accc
added option to control soundness of encodings more precisely, for evaluation purposes
blanchet
parents:
44393
diff
changeset
|
413 |
especially if types are mangled in names. This makes a difference for some |
|
20bd9f90accc
added option to control soundness of encodings more precisely, for evaluation purposes
blanchet
parents:
44393
diff
changeset
|
414 |
provers (e.g., E). For these reason, short names are enabled by default. *) |
|
44592
54906b0337ab
flip logic of boolean option so it's off by default
blanchet
parents:
44586
diff
changeset
|
415 |
val atp_full_names = |
|
54906b0337ab
flip logic of boolean option so it's off by default
blanchet
parents:
44586
diff
changeset
|
416 |
Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
|
|
43092
93ec303e1917
more work on new metis that exploits the powerful new type encodings
blanchet
parents:
43085
diff
changeset
|
417 |
|
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
418 |
val smt_triggers = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
419 |
Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
|
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
420 |
val smt_weights = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
421 |
Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
|
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
422 |
val smt_weight_min_facts = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
423 |
Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
|
|
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
424 |
|
|
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
425 |
(* FUDGE *) |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
426 |
val smt_min_weight = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
427 |
Attrib.setup_config_int @{binding sledgehammer_smt_min_weight} (K 0)
|
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
428 |
val smt_max_weight = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
429 |
Attrib.setup_config_int @{binding sledgehammer_smt_max_weight} (K 10)
|
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
430 |
val smt_max_weight_index = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
431 |
Attrib.setup_config_int @{binding sledgehammer_smt_max_weight_index} (K 200)
|
|
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
432 |
val smt_weight_curve = Unsynchronized.ref (fn x : int => x * x) |
|
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
433 |
|
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
434 |
fun smt_fact_weight ctxt j num_facts = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
435 |
if Config.get ctxt smt_weights andalso |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
436 |
num_facts >= Config.get ctxt smt_weight_min_facts then |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
437 |
let |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
438 |
val min = Config.get ctxt smt_min_weight |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
439 |
val max = Config.get ctxt smt_max_weight |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
440 |
val max_index = Config.get ctxt smt_max_weight_index |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
441 |
val curve = !smt_weight_curve |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
442 |
in |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
443 |
SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1)) |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
444 |
div curve max_index) |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
445 |
end |
|
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
446 |
else |
|
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
447 |
NONE |
|
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
448 |
|
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
449 |
fun weight_smt_fact ctxt num_facts ((info, th), j) = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
450 |
let val thy = Proof_Context.theory_of ctxt in |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
451 |
(info, (smt_fact_weight ctxt j num_facts, th |> Thm.transfer thy)) |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
452 |
end |
| 38023 | 453 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
454 |
fun overlord_file_location_of_prover prover = |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
455 |
(getenv "ISABELLE_HOME_USER", "prob_" ^ prover) |
|
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
456 |
|
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
457 |
fun proof_banner mode name = |
| 43033 | 458 |
case mode of |
459 |
Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof"
|
|
460 |
| Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
|
|
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
461 |
| _ => "Try this" |
| 43033 | 462 |
|
|
45561
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
463 |
fun bunch_of_reconstructors needs_full_types lam_trans = |
| 48800 | 464 |
if needs_full_types then |
| 48802 | 465 |
[Metis (full_type_enc, lam_trans false), |
466 |
Metis (really_full_type_enc, lam_trans false), |
|
467 |
Metis (full_type_enc, lam_trans true), |
|
468 |
Metis (really_full_type_enc, lam_trans true), |
|
469 |
SMT] |
|
470 |
else |
|
| 48800 | 471 |
[Metis (partial_type_enc, lam_trans false), |
472 |
Metis (full_type_enc, lam_trans false), |
|
473 |
Metis (no_typesN, lam_trans true), |
|
474 |
Metis (really_full_type_enc, lam_trans true), |
|
475 |
SMT] |
|
|
45561
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
476 |
|
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
477 |
fun extract_reconstructor ({type_enc, lam_trans, ...} : params)
|
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
478 |
(Metis (type_enc', lam_trans')) = |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
479 |
let |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
480 |
val override_params = |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
481 |
(if is_none type_enc andalso type_enc' = hd partial_type_encs then |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
482 |
[] |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
483 |
else |
|
45566
da05ce2de5a8
better threading of type encodings between Sledgehammer and "metis"
blanchet
parents:
45561
diff
changeset
|
484 |
[("type_enc", [hd (unalias_type_enc type_enc')])]) @
|
|
45561
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
485 |
(if is_none lam_trans andalso lam_trans' = metis_default_lam_trans then |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
486 |
[] |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
487 |
else |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
488 |
[("lam_trans", [lam_trans'])])
|
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
489 |
in (metisN, override_params) end |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
490 |
| extract_reconstructor _ SMT = (smtN, []) |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
491 |
|
| 43033 | 492 |
(* based on "Mirabelle.can_apply" and generalized *) |
|
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
493 |
fun timed_apply timeout tac state i = |
| 43033 | 494 |
let |
495 |
val {context = ctxt, facts, goal} = Proof.goal state
|
|
496 |
val full_tac = Method.insert_tac facts i THEN tac ctxt i |
|
| 50557 | 497 |
in time_limit timeout (try (Seq.pull o full_tac)) goal end |
| 43033 | 498 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
499 |
fun tac_of_reconstructor (Metis (type_enc, lam_trans)) = |
| 45521 | 500 |
metis_tac [type_enc] lam_trans |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
501 |
| tac_of_reconstructor SMT = SMT_Solver.smt_tac |
|
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
502 |
|
| 45520 | 503 |
fun timed_reconstructor reconstr debug timeout ths = |
|
44651
5d6a11e166cf
renamed "Metis_Tactics" to "Metis_Tactic", now that there is only one Metis tactic ("metisFT" is legacy)
blanchet
parents:
44649
diff
changeset
|
504 |
(Config.put Metis_Tactic.verbose debug |
| 45557 | 505 |
#> Config.put SMT_Config.verbose debug |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
506 |
#> (fn ctxt => tac_of_reconstructor reconstr ctxt ths)) |
|
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
507 |
|> timed_apply timeout |
| 43033 | 508 |
|
| 48798 | 509 |
fun is_fact_chained ((_, (sc, _)), _) = sc = Chained |
510 |
||
511 |
fun filter_used_facts keep_chained used = |
|
512 |
filter ((member (op =) used o fst) orf |
|
513 |
(if keep_chained then is_fact_chained else K false)) |
|
| 43033 | 514 |
|
|
45781
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
blanchet
parents:
45707
diff
changeset
|
515 |
fun play_one_line_proof mode debug verbose timeout pairs state i preferred |
| 45520 | 516 |
reconstrs = |
|
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
517 |
let |
| 45520 | 518 |
fun get_preferred reconstrs = |
519 |
if member (op =) reconstrs preferred then preferred |
|
520 |
else List.last reconstrs |
|
|
43034
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
blanchet
parents:
43033
diff
changeset
|
521 |
in |
| 50557 | 522 |
if timeout = SOME Time.zeroTime then |
| 45520 | 523 |
Trust_Playable (get_preferred reconstrs, NONE) |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
524 |
else |
| 50557 | 525 |
let |
526 |
val _ = |
|
527 |
if mode = Minimize then Output.urgent_message "Preplaying proof..." |
|
528 |
else () |
|
529 |
val ths = pairs |> sort_wrt (fst o fst) |> map snd |
|
530 |
fun play [] [] = Failed_to_Play (get_preferred reconstrs) |
|
531 |
| play timed_outs [] = |
|
532 |
Trust_Playable (get_preferred timed_outs, timeout) |
|
533 |
| play timed_out (reconstr :: reconstrs) = |
|
534 |
let |
|
535 |
val _ = |
|
536 |
if verbose then |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
537 |
"Trying \"" ^ string_of_reconstructor reconstr ^ "\"" ^ |
| 50557 | 538 |
(case timeout of |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
539 |
SOME timeout => " for " ^ string_of_time timeout |
| 50557 | 540 |
| NONE => "") ^ "..." |
541 |
|> Output.urgent_message |
|
542 |
else |
|
543 |
() |
|
544 |
val timer = Timer.startRealTimer () |
|
545 |
in |
|
546 |
case timed_reconstructor reconstr debug timeout ths state i of |
|
547 |
SOME (SOME _) => Played (reconstr, Timer.checkRealTimer timer) |
|
548 |
| _ => play timed_out reconstrs |
|
549 |
end |
|
550 |
handle TimeLimit.TimeOut => play (reconstr :: timed_out) reconstrs |
|
551 |
in play [] reconstrs end |
|
| 43033 | 552 |
end |
553 |
||
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41259
diff
changeset
|
554 |
|
|
41256
0e7d45cc005f
put the SMT weights back where they belong, so that they're also used by Mirabelle
blanchet
parents:
41242
diff
changeset
|
555 |
(* generic TPTP-based ATPs *) |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
556 |
|
|
42730
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
557 |
(* Too general means, positive equality literal with a variable X as one |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
558 |
operand, when X does not occur properly in the other operand. This rules out |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
559 |
clearly inconsistent facts such as X = a | X = b, though it by no means |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
560 |
guarantees soundness. *) |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
561 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
562 |
fun get_facts_of_filter _ [(_, facts)] = facts |
|
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
563 |
| get_facts_of_filter fact_filter factss = |
| 51013 | 564 |
case AList.lookup (op =) factss fact_filter of |
565 |
SOME facts => facts |
|
566 |
| NONE => snd (hd factss) |
|
567 |
||
|
42730
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
568 |
(* Unwanted equalities are those between a (bound or schematic) variable that |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
569 |
does not properly occur in the second operand. *) |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
570 |
val is_exhaustive_finite = |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
571 |
let |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
572 |
fun is_bad_equal (Var z) t = |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
573 |
not (exists_subterm (fn Var z' => z = z' | _ => false) t) |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
574 |
| is_bad_equal (Bound j) t = not (loose_bvar1 (t, j)) |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
575 |
| is_bad_equal _ _ = false |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
576 |
fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1 |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
577 |
fun do_formula pos t = |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
578 |
case (pos, t) of |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
579 |
(_, @{const Trueprop} $ t1) => do_formula pos t1
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
580 |
| (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
581 |
do_formula pos t' |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
582 |
| (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
583 |
do_formula pos t' |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
584 |
| (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
585 |
do_formula pos t' |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
586 |
| (_, @{const "==>"} $ t1 $ t2) =>
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
587 |
do_formula (not pos) t1 andalso |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
588 |
(t2 = @{prop False} orelse do_formula pos t2)
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
589 |
| (_, @{const HOL.implies} $ t1 $ t2) =>
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
590 |
do_formula (not pos) t1 andalso |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
591 |
(t2 = @{const False} orelse do_formula pos t2)
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
592 |
| (_, @{const Not} $ t1) => do_formula (not pos) t1
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
593 |
| (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
594 |
| (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
595 |
| (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
596 |
| (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
597 |
| _ => false |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
598 |
in do_formula true end |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
599 |
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
600 |
fun has_bound_or_var_of_type pred = |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
601 |
exists_subterm (fn Var (_, T as Type _) => pred T |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
602 |
| Abs (_, T as Type _, _) => pred T |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
603 |
| _ => false) |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
604 |
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
605 |
(* Facts are forbidden to contain variables of these types. The typical reason |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
606 |
is that they lead to unsoundness. Note that "unit" satisfies numerous |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
607 |
equations like "?x = ()". The resulting clauses will have no type constraint, |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
608 |
yielding false proofs. Even "bool" leads to many unsound proofs, though only |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
609 |
for higher-order problems. *) |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
610 |
|
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
611 |
(* Facts containing variables of type "unit" or "bool" or of the form |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
612 |
"ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
613 |
are omitted. *) |
|
42944
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
614 |
fun is_dangerous_prop ctxt = |
|
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents:
42943
diff
changeset
|
615 |
transform_elim_prop |
| 44393 | 616 |
#> (has_bound_or_var_of_type (is_type_surely_finite ctxt) orf |
|
42730
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
617 |
is_exhaustive_finite) |
|
d6db5a815477
improve detection of quantifications over dangerous types by leveraging "is_type_surely_finite" predicate and added "prop" to the list of surely finite types
blanchet
parents:
42729
diff
changeset
|
618 |
|
| 39492 | 619 |
(* Important messages are important but not so important that users want to see |
620 |
them each time. *) |
|
| 44649 | 621 |
val atp_important_message_keep_quotient = 25 |
| 39492 | 622 |
|
|
44416
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
blanchet
parents:
44397
diff
changeset
|
623 |
fun choose_type_enc soundness best_type_enc format = |
|
44397
06375952f1fa
cleaner handling of polymorphic monotonicity inference
blanchet
parents:
44394
diff
changeset
|
624 |
the_default best_type_enc |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
625 |
#> type_enc_of_string soundness |
|
44416
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
blanchet
parents:
44397
diff
changeset
|
626 |
#> adjust_type_enc format |
|
42548
ea2a28b1938f
make sure the minimizer monomorphizes when it should
blanchet
parents:
42544
diff
changeset
|
627 |
|
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
628 |
fun isar_proof_reconstructor ctxt name = |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
629 |
let val thy = Proof_Context.theory_of ctxt in |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
630 |
if is_atp thy name then name |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
631 |
else remotify_prover_if_not_installed ctxt eN |> the_default name |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
632 |
end |
| 43051 | 633 |
|
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
634 |
(* See the analogous logic in the function "maybe_minimize" in |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
635 |
"sledgehammer_minimize.ML". *) |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
636 |
fun choose_minimize_command ctxt (params as {isar_proofs, ...}) minimize_command
|
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
637 |
name preplay = |
| 45520 | 638 |
let |
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
639 |
val maybe_isar_name = |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
640 |
name |> isar_proofs = SOME true ? isar_proof_reconstructor ctxt |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
641 |
val (min_name, override_params) = |
| 45520 | 642 |
case preplay of |
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
643 |
Played (reconstr, _) => |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
644 |
if isar_proofs = SOME true then (maybe_isar_name, []) |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
645 |
else extract_reconstructor params reconstr |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
646 |
| _ => (maybe_isar_name, []) |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
647 |
in minimize_command override_params min_name end |
| 43051 | 648 |
|
|
52034
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
649 |
fun repair_legacy_monomorph_context max_iters best_max_iters max_new_instances |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
650 |
best_max_new_instances = |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
651 |
Config.put Legacy_Monomorph.max_rounds |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
652 |
(max_iters |> the_default best_max_iters) |
| 51575 | 653 |
#> Config.put Legacy_Monomorph.max_new_instances |
|
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47946
diff
changeset
|
654 |
(max_new_instances |> the_default best_max_new_instances) |
| 51575 | 655 |
#> Config.put Legacy_Monomorph.keep_partial_instances false |
| 43226 | 656 |
|
|
52034
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
657 |
fun repair_monomorph_context max_iters best_max_iters max_new_instances |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
658 |
best_max_new_instances = |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
659 |
Config.put Monomorph.max_rounds (max_iters |> the_default best_max_iters) |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
660 |
#> Config.put Monomorph.max_new_instances |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
661 |
(max_new_instances |> the_default best_max_new_instances) |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
662 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
663 |
fun suffix_of_mode Auto_Try = "_try" |
|
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
664 |
| suffix_of_mode Try = "_try" |
|
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
665 |
| suffix_of_mode Normal = "" |
|
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
666 |
| suffix_of_mode MaSh = "" |
|
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
667 |
| suffix_of_mode Auto_Minimize = "_min" |
|
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
668 |
| suffix_of_mode Minimize = "_min" |
|
44509
369e8c28a61a
added a component in generated file names reflecting whether the minimizer is used -- needed for evaluation to keep these files separated from the main problem files
blanchet
parents:
44423
diff
changeset
|
669 |
|
|
44423
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents:
44416
diff
changeset
|
670 |
(* Give the ATPs some slack before interrupting them the hard way. "z3_tptp" on |
|
43631
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
blanchet
parents:
43626
diff
changeset
|
671 |
Linux appears to be the only ATP that does not honor its time limit. *) |
| 43690 | 672 |
val atp_timeout_slack = seconds 1.0 |
|
43631
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
blanchet
parents:
43626
diff
changeset
|
673 |
|
|
48331
f190a6dbb29b
make the monomorphizer more predictable by making the cutoff independent on the number of facts
blanchet
parents:
48321
diff
changeset
|
674 |
val mono_max_privileged_facts = 10 |
|
f190a6dbb29b
make the monomorphizer more predictable by making the cutoff independent on the number of facts
blanchet
parents:
48321
diff
changeset
|
675 |
|
|
51186
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
676 |
(* For low values of "max_facts", this fudge value ensures that most slices are |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
677 |
invoked with a nontrivial amount of facts. *) |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
678 |
val max_fact_factor_fudge = 5 |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
679 |
|
| 43021 | 680 |
fun run_atp mode name |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
681 |
({exec, arguments, proof_delims, known_failures, prem_role, best_slices,
|
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
682 |
best_max_mono_iters, best_max_new_mono_instances, ...} : atp_config) |
| 46301 | 683 |
(params as {debug, verbose, overlord, type_enc, strict, lam_trans,
|
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
684 |
uncurried_aliases, fact_filter, max_facts, max_mono_iters, |
|
51130
76d68444cd59
renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents:
51024
diff
changeset
|
685 |
max_new_mono_instances, isar_proofs, isar_compress, |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52555
diff
changeset
|
686 |
isar_compress_degree, merge_timeout_slack, |
| 51879 | 687 |
slice, timeout, preplay_timeout, preplay_trace, ...}) |
| 43037 | 688 |
minimize_command |
| 51013 | 689 |
({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
|
| 38023 | 690 |
let |
|
42182
a630978fc967
start monomorphization process with subgoal, not entire goal, to avoid needless instances (and only print monomorphization messages in debug mode)
blanchet
parents:
42181
diff
changeset
|
691 |
val thy = Proof.theory_of state |
| 39318 | 692 |
val ctxt = Proof.context_of state |
|
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47934
diff
changeset
|
693 |
val atp_mode = |
| 48143 | 694 |
if Config.get ctxt completish then Sledgehammer_Completish |
|
47946
33afcfad3f8d
add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents:
47934
diff
changeset
|
695 |
else Sledgehammer |
|
52196
2281f33e8da6
redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents:
52150
diff
changeset
|
696 |
val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt |
|
41159
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
697 |
val (dest_dir, problem_prefix) = |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
698 |
if overlord then overlord_file_location_of_prover name |
|
41159
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
699 |
else (Config.get ctxt dest_dir, Config.get ctxt problem_prefix) |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
700 |
val problem_file_name = |
|
41159
1e12d6495423
honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents:
41152
diff
changeset
|
701 |
Path.basic (problem_prefix ^ (if overlord then "" else serial_string ()) ^ |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
702 |
suffix_of_mode mode ^ "_" ^ string_of_int subgoal) |
|
48656
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48532
diff
changeset
|
703 |
val prob_path = |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
704 |
if dest_dir = "" then |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
705 |
File.tmp_path problem_file_name |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
706 |
else if File.exists (Path.explode dest_dir) then |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
707 |
Path.append (Path.explode dest_dir) problem_file_name |
|
39003
c2aebd79981f
run relevance filter in a thread, to avoid blocking
blanchet
parents:
39000
diff
changeset
|
708 |
else |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39494
diff
changeset
|
709 |
error ("No such directory: " ^ quote dest_dir ^ ".")
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
710 |
val command0 = |
|
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47038
diff
changeset
|
711 |
case find_first (fn var => getenv var <> "") (fst exec) of |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
712 |
SOME var => |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
713 |
let |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
714 |
val pref = getenv var ^ "/" |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
715 |
val paths = map (Path.explode o prefix pref) (snd exec) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
716 |
in |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
717 |
case find_first File.exists paths of |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
718 |
SOME path => path |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
719 |
| NONE => error ("Bad executable: " ^ Path.print (hd paths) ^ ".")
|
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
720 |
end |
|
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47038
diff
changeset
|
721 |
| NONE => error ("The environment variable " ^ quote (hd (fst exec)) ^
|
|
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47038
diff
changeset
|
722 |
" is not set.") |
| 38023 | 723 |
fun split_time s = |
724 |
let |
|
| 42448 | 725 |
val split = String.tokens (fn c => str c = "\n") |
|
47737
63c939dcd055
made "split_last" more robust in the face of obscure low-level errors
blanchet
parents:
47606
diff
changeset
|
726 |
val (output, t) = |
|
63c939dcd055
made "split_last" more robust in the face of obscure low-level errors
blanchet
parents:
47606
diff
changeset
|
727 |
s |> split |> (try split_last #> the_default ([], "0")) |
|
63c939dcd055
made "split_last" more robust in the face of obscure low-level errors
blanchet
parents:
47606
diff
changeset
|
728 |
|>> cat_lines |
| 42448 | 729 |
fun as_num f = f >> (fst o read_int) |
730 |
val num = as_num (Scan.many1 Symbol.is_ascii_digit) |
|
731 |
val digit = Scan.one Symbol.is_ascii_digit |
|
732 |
val num3 = as_num (digit ::: digit ::: (digit >> single)) |
|
733 |
val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b) |
|
| 45381 | 734 |
val as_time = |
735 |
raw_explode #> Scan.read Symbol.stopper time #> the_default 0 |
|
|
47737
63c939dcd055
made "split_last" more robust in the face of obscure low-level errors
blanchet
parents:
47606
diff
changeset
|
736 |
in (output, as_time t |> Time.fromMilliseconds) end |
|
48656
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48532
diff
changeset
|
737 |
fun run () = |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
738 |
let |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
739 |
(* If slicing is disabled, we expand the last slice to fill the entire |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
740 |
time available. *) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
741 |
val actual_slices = get_slices slice (best_slices ctxt) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
742 |
val num_actual_slices = length actual_slices |
|
51186
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
743 |
val max_fact_factor = |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
744 |
case max_facts of |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
745 |
NONE => 1.0 |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
746 |
| SOME max => |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
747 |
Real.fromInt max |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
748 |
/ Real.fromInt (fold (Integer.max o slice_max_facts) |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
749 |
actual_slices 0) |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
750 |
fun monomorphize_facts facts = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
751 |
let |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
752 |
val ctxt = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
753 |
ctxt |
|
52034
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
754 |
|> repair_legacy_monomorph_context max_mono_iters |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
755 |
best_max_mono_iters max_new_mono_instances |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
756 |
best_max_new_mono_instances |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
757 |
(* pseudo-theorem involving the same constants as the subgoal *) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
758 |
val subgoal_th = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
759 |
Logic.list_implies (hyp_ts, concl_t) |> Skip_Proof.make_thm thy |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
760 |
val rths = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
761 |
facts |> chop mono_max_privileged_facts |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
762 |
|>> map (pair 1 o snd) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
763 |
||> map (pair 2 o snd) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
764 |
|> op @ |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
765 |
|> cons (0, subgoal_th) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
766 |
in |
| 51575 | 767 |
Legacy_Monomorph.monomorph atp_schematic_consts_of rths ctxt |> fst |> tl |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
768 |
|> curry ListPair.zip (map fst facts) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
769 |
|> maps (fn (name, rths) => |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
770 |
map (pair name o zero_var_indexes o snd) rths) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
771 |
end |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
772 |
fun run_slice time_left (cache_key, cache_value) |
|
48716
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48656
diff
changeset
|
773 |
(slice, (time_frac, |
| 51011 | 774 |
(key as ((best_max_facts, best_fact_filter), format, |
775 |
best_type_enc, best_lam_trans, |
|
776 |
best_uncurried_aliases), |
|
|
48716
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48656
diff
changeset
|
777 |
extra))) = |
| 38032 | 778 |
let |
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
779 |
val effective_fact_filter = |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
780 |
fact_filter |> the_default best_fact_filter |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
781 |
val facts = get_facts_of_filter effective_fact_filter factss |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
782 |
val num_facts = |
|
51186
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
783 |
Real.ceil (max_fact_factor * Real.fromInt best_max_facts) + |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
784 |
max_fact_factor_fudge |
|
c8721406511a
interpret "max_facts" argument in a slice-dependent fashion, instead of forcing the same number of facts to all slices
blanchet
parents:
51181
diff
changeset
|
785 |
|> Integer.min (length facts) |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
786 |
val soundness = if strict then Strict else Non_Strict |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
787 |
val type_enc = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
788 |
type_enc |> choose_type_enc soundness best_type_enc format |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
789 |
val sound = is_type_enc_sound type_enc |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
790 |
val real_ms = Real.fromInt o Time.toMilliseconds |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
791 |
val slice_timeout = |
| 50557 | 792 |
case time_left of |
793 |
SOME time_left => |
|
794 |
((real_ms time_left |
|
795 |
|> (if slice < num_actual_slices - 1 then |
|
796 |
curry Real.min (time_frac * real_ms (the timeout)) |
|
797 |
else |
|
798 |
I)) |
|
799 |
* 0.001) |
|
800 |
|> seconds |> SOME |
|
801 |
| NONE => NONE |
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
802 |
val generous_slice_timeout = |
|
50558
a719106124d8
avoid creating nested threads for MaSh -- this seems to cause thread creation failures for machines with dozens of cores (unclear yet if that's really the issue)
blanchet
parents:
50557
diff
changeset
|
803 |
if mode = MaSh then NONE |
|
a719106124d8
avoid creating nested threads for MaSh -- this seems to cause thread creation failures for machines with dozens of cores (unclear yet if that's really the issue)
blanchet
parents:
50557
diff
changeset
|
804 |
else Option.map (curry Time.+ atp_timeout_slack) slice_timeout |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
805 |
val _ = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
806 |
if debug then |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
807 |
quote name ^ " slice #" ^ string_of_int (slice + 1) ^ |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
808 |
" with " ^ string_of_int num_facts ^ " fact" ^ |
| 50557 | 809 |
plural_s num_facts ^ |
810 |
(case slice_timeout of |
|
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
811 |
SOME timeout => " for " ^ string_of_time timeout |
| 50557 | 812 |
| NONE => "") ^ "..." |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
813 |
|> Output.urgent_message |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
814 |
else |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
815 |
() |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
816 |
val readable_names = not (Config.get ctxt atp_full_names) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
817 |
val lam_trans = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
818 |
case lam_trans of |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
819 |
SOME s => s |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
820 |
| NONE => best_lam_trans |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
821 |
val uncurried_aliases = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
822 |
case uncurried_aliases of |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
823 |
SOME b => b |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
824 |
| NONE => best_uncurried_aliases |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
825 |
val value as (atp_problem, _, fact_names, _, _) = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
826 |
if cache_key = SOME key then |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
827 |
cache_value |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
828 |
else |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
829 |
facts |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
830 |
|> not sound |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
831 |
? filter_out (is_dangerous_prop ctxt o prop_of o snd) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
832 |
|> take num_facts |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
833 |
|> not (is_type_enc_polymorphic type_enc) ? monomorphize_facts |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
834 |
|> map (apsnd prop_of) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
835 |
|> prepare_atp_problem ctxt format prem_role type_enc atp_mode |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
836 |
lam_trans uncurried_aliases |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
837 |
readable_names true hyp_ts concl_t |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
838 |
fun sel_weights () = atp_problem_selection_weights atp_problem |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
839 |
fun ord_info () = atp_problem_term_order_info atp_problem |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
840 |
val ord = effective_term_order ctxt name |
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
841 |
val full_proof = |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
842 |
debug orelse (isar_proofs |> the_default (mode = Minimize)) |
| 50927 | 843 |
val args = |
844 |
arguments ctxt full_proof extra |
|
845 |
(slice_timeout |> the_default one_day) |
|
846 |
(File.shell_path prob_path) (ord, ord_info, sel_weights) |
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
847 |
val command = |
| 50927 | 848 |
"(exec 2>&1; " ^ File.shell_path command0 ^ " " ^ args ^ " " ^ ")" |
|
48532
c0f44941e674
Sledgehammer already has its own ways of reporting and recovering from crashes in external provers -- no need to additionally print scores of warnings (cf. 4b0daca2bf88)
blanchet
parents:
48392
diff
changeset
|
849 |
|> enclose "TIMEFORMAT='%3R'; { time " " ; }"
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
850 |
val _ = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
851 |
atp_problem |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
852 |
|> lines_of_atp_problem format ord ord_info |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
853 |
|> cons ("% " ^ command ^ "\n")
|
|
48656
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48532
diff
changeset
|
854 |
|> File.write_list prob_path |
|
52034
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
855 |
val ((output, run_time), (atp_proof, outcome)) = |
| 50557 | 856 |
time_limit generous_slice_timeout Isabelle_System.bash_output |
857 |
command |
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
858 |
|>> (if overlord then |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
859 |
prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
|
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
860 |
else |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
861 |
I) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
862 |
|> fst |> split_time |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
863 |
|> (fn accum as (output, _) => |
|
52034
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
864 |
(accum, |
|
48716
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48656
diff
changeset
|
865 |
extract_tstplike_proof_and_outcome verbose proof_delims |
|
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48656
diff
changeset
|
866 |
known_failures output |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
867 |
|>> atp_proof_of_tstplike_proof atp_problem |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
868 |
handle UNRECOGNIZED_ATP_PROOF () => |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
869 |
([], SOME ProofIncomplete))) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
870 |
handle TimeLimit.TimeOut => |
|
52034
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
871 |
(("", the slice_timeout), ([], SOME TimedOut))
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
872 |
val outcome = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
873 |
case outcome of |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
874 |
NONE => |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
875 |
(case used_facts_in_unsound_atp_proof ctxt fact_names atp_proof |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
876 |
|> Option.map (sort string_ord) of |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
877 |
SOME facts => |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
878 |
let |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
879 |
val failure = |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
880 |
UnsoundProof (is_type_enc_sound type_enc, facts) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
881 |
in |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
882 |
if debug then (warning (string_of_failure failure); NONE) |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
883 |
else SOME failure |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
884 |
end |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
885 |
| NONE => NONE) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
886 |
| _ => outcome |
| 51013 | 887 |
in |
888 |
((SOME key, value), (output, run_time, facts, atp_proof, outcome)) |
|
889 |
end |
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
890 |
val timer = Timer.startRealTimer () |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
891 |
fun maybe_run_slice slice |
| 51013 | 892 |
(result as (cache, (_, run_time0, _, _, SOME _))) = |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
893 |
let |
| 50557 | 894 |
val time_left = |
895 |
Option.map |
|
896 |
(fn timeout => Time.- (timeout, Timer.checkRealTimer timer)) |
|
897 |
timeout |
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
898 |
in |
| 50557 | 899 |
if time_left <> NONE andalso |
900 |
Time.<= (the time_left, Time.zeroTime) then |
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
901 |
result |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
902 |
else |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
903 |
run_slice time_left cache slice |
| 51013 | 904 |
|> (fn (cache, (output, run_time, used_from, atp_proof, |
905 |
outcome)) => |
|
906 |
(cache, (output, Time.+ (run_time0, run_time), used_from, |
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
907 |
atp_proof, outcome))) |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
908 |
end |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
909 |
| maybe_run_slice _ result = result |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
910 |
in |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
911 |
((NONE, ([], Symtab.empty, Vector.fromList [], [], Symtab.empty)), |
| 51013 | 912 |
("", Time.zeroTime, [], [], SOME InternalError))
|
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
913 |
|> fold maybe_run_slice actual_slices |
|
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
914 |
end |
| 38023 | 915 |
(* If the problem file has not been exported, remove it; otherwise, export |
916 |
the proof file too. *) |
|
|
48656
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48532
diff
changeset
|
917 |
fun clean_up () = |
|
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48532
diff
changeset
|
918 |
if dest_dir = "" then (try File.rm prob_path; ()) else () |
| 51013 | 919 |
fun export (_, (output, _, _, _, _)) = |
|
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48331
diff
changeset
|
920 |
if dest_dir = "" then () |
|
48656
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48532
diff
changeset
|
921 |
else File.write (Path.explode (Path.implode prob_path ^ "_proof")) output |
|
52150
41c885784e04
handle lambda-lifted problems in Isar construction code
blanchet
parents:
52125
diff
changeset
|
922 |
val ((_, (_, pool, fact_names, lifted, sym_tab)), |
| 51013 | 923 |
(output, run_time, used_from, atp_proof, outcome)) = |
|
48656
5caa414ce9a2
cleaner temporary file cleanup for MaSh, based on tried-and-trusted code
blanchet
parents:
48532
diff
changeset
|
924 |
with_cleanup clean_up run () |> tap export |
| 39492 | 925 |
val important_message = |
| 43021 | 926 |
if mode = Normal andalso |
| 42609 | 927 |
random_range 0 (atp_important_message_keep_quotient - 1) = 0 then |
| 39492 | 928 |
extract_important_message output |
929 |
else |
|
930 |
"" |
|
|
43261
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
931 |
val (used_facts, preplay, message, message_tail) = |
| 38023 | 932 |
case outcome of |
933 |
NONE => |
|
| 43033 | 934 |
let |
| 45551 | 935 |
val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof |
| 45590 | 936 |
val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof |
| 45521 | 937 |
val reconstrs = |
|
45554
09ad83de849c
don't pass "lam_lifted" option to "metis" unless there's a good reason
blanchet
parents:
45553
diff
changeset
|
938 |
bunch_of_reconstructors needs_full_types |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
939 |
(lam_trans_of_atp_proof atp_proof |
| 46405 | 940 |
o (fn desperate => if desperate then hide_lamsN |
941 |
else metis_default_lam_trans)) |
|
| 43033 | 942 |
in |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
943 |
(used_facts, |
|
50669
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
944 |
Lazy.lazy (fn () => |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
945 |
let |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
946 |
val used_pairs = |
| 51013 | 947 |
used_from |> filter_used_facts false used_facts |
|
50669
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
948 |
in |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
949 |
play_one_line_proof mode debug verbose preplay_timeout |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
950 |
used_pairs state subgoal (hd reconstrs) reconstrs |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
951 |
end), |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
952 |
fn preplay => |
|
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
953 |
let |
| 49921 | 954 |
val _ = |
955 |
if verbose then |
|
| 51232 | 956 |
Output.urgent_message "Generating proof text..." |
| 49921 | 957 |
else |
958 |
() |
|
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
959 |
val isar_params = |
| 51879 | 960 |
(debug, verbose, preplay_timeout, preplay_trace, isar_compress, |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52555
diff
changeset
|
961 |
isar_compress_degree, merge_timeout_slack, |
|
52150
41c885784e04
handle lambda-lifted problems in Isar construction code
blanchet
parents:
52125
diff
changeset
|
962 |
pool, fact_names, lifted, sym_tab, atp_proof, goal) |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
963 |
val one_line_params = |
|
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
964 |
(preplay, proof_banner mode name, used_facts, |
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
965 |
choose_minimize_command ctxt params minimize_command name |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
966 |
preplay, |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
967 |
subgoal, subgoal_count) |
| 48799 | 968 |
val num_chained = length (#facts (Proof.goal state)) |
969 |
in |
|
|
49918
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents:
49914
diff
changeset
|
970 |
proof_text ctxt isar_proofs isar_params num_chained |
| 48799 | 971 |
one_line_params |
972 |
end, |
|
|
43261
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
973 |
(if verbose then |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
974 |
"\nATP real CPU time: " ^ string_of_time run_time ^ "." |
|
43261
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
975 |
else |
|
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
976 |
"") ^ |
|
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
977 |
(if important_message <> "" then |
|
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
978 |
"\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ |
|
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
979 |
important_message |
|
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
980 |
else |
|
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
981 |
"")) |
| 43033 | 982 |
end |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
983 |
| SOME failure => |
|
50669
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
984 |
([], Lazy.value (Failed_to_Play plain_metis), |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
985 |
fn _ => string_of_failure failure, "") |
| 38023 | 986 |
in |
| 51013 | 987 |
{outcome = outcome, used_facts = used_facts, used_from = used_from,
|
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
988 |
run_time = run_time, preplay = preplay, message = message, |
|
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
989 |
message_tail = message_tail} |
| 38023 | 990 |
end |
991 |
||
| 51014 | 992 |
fun rotate_one (x :: xs) = xs @ [x] |
993 |
||
| 40669 | 994 |
(* "SMT_Failure.Abnormal_Termination" carries the solver's return code. Until |
995 |
these are sorted out properly in the SMT module, we have to interpret these |
|
996 |
ourselves. *) |
|
|
40684
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
997 |
val remote_smt_failures = |
|
43631
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
blanchet
parents:
43626
diff
changeset
|
998 |
[(2, NoLibwwwPerl), |
|
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
blanchet
parents:
43626
diff
changeset
|
999 |
(22, CantConnect)] |
|
40684
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
1000 |
val z3_failures = |
| 41236 | 1001 |
[(101, OutOfResources), |
1002 |
(103, MalformedInput), |
|
| 50667 | 1003 |
(110, MalformedInput), |
1004 |
(112, TimedOut)] |
|
|
40684
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
1005 |
val unix_failures = |
| 48797 | 1006 |
[(138, Crashed), |
1007 |
(139, Crashed)] |
|
|
43631
4144d7b4ec77
enforce hard timeout on ATPs (esp. "z3_atp" on Linux) + remove obsolete failure codes
blanchet
parents:
43626
diff
changeset
|
1008 |
val smt_failures = remote_smt_failures @ z3_failures @ unix_failures |
| 40555 | 1009 |
|
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1010 |
fun failure_of_smt_failure (SMT_Failure.Counterexample {is_real_cex, ...}) =
|
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1011 |
if is_real_cex then Unprovable else GaveUp |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1012 |
| failure_of_smt_failure SMT_Failure.Time_Out = TimedOut |
|
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1013 |
| failure_of_smt_failure (SMT_Failure.Abnormal_Termination code) = |
| 41222 | 1014 |
(case AList.lookup (op =) smt_failures code of |
|
40684
c7ba327eb58c
more precise error handling in Sledgehammer/SMT for Z3 and remote SMT provers
blanchet
parents:
40669
diff
changeset
|
1015 |
SOME failure => failure |
|
41259
13972ced98d9
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
blanchet
parents:
41256
diff
changeset
|
1016 |
| NONE => UnknownError ("Abnormal termination with exit code " ^
|
|
13972ced98d9
more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
blanchet
parents:
41256
diff
changeset
|
1017 |
string_of_int code ^ ".")) |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1018 |
| failure_of_smt_failure SMT_Failure.Out_Of_Memory = OutOfResources |
|
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1019 |
| failure_of_smt_failure (SMT_Failure.Other_Failure s) = UnknownError s |
|
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1020 |
|
| 40698 | 1021 |
(* FUDGE *) |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
1022 |
val smt_max_slices = |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
1023 |
Attrib.setup_config_int @{binding sledgehammer_smt_max_slices} (K 8)
|
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
1024 |
val smt_slice_fact_frac = |
| 51014 | 1025 |
Attrib.setup_config_real @{binding sledgehammer_smt_slice_fact_frac}
|
1026 |
(K 0.667) |
|
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
1027 |
val smt_slice_time_frac = |
| 51014 | 1028 |
Attrib.setup_config_real @{binding sledgehammer_smt_slice_time_frac} (K 0.333)
|
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
1029 |
val smt_slice_min_secs = |
| 51014 | 1030 |
Attrib.setup_config_int @{binding sledgehammer_smt_slice_min_secs} (K 3)
|
|
40409
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents:
40370
diff
changeset
|
1031 |
|
| 50759 | 1032 |
fun smt_filter_loop name |
|
42724
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents:
42723
diff
changeset
|
1033 |
({debug, verbose, overlord, max_mono_iters,
|
| 45706 | 1034 |
max_new_mono_instances, timeout, slice, ...} : params) |
|
50486
d5dc28fafd9d
made MaSh evaluation driver work with SMT solvers
blanchet
parents:
50020
diff
changeset
|
1035 |
state goal i = |
|
40409
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents:
40370
diff
changeset
|
1036 |
let |
| 50759 | 1037 |
fun repair_context ctxt = |
1038 |
ctxt |> select_smt_solver name |
|
1039 |
|> Config.put SMT_Config.verbose debug |
|
1040 |
|> (if overlord then |
|
1041 |
Config.put SMT_Config.debug_files |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
1042 |
(overlord_file_location_of_prover name |
| 50759 | 1043 |
|> (fn (path, name) => path ^ "/" ^ name)) |
1044 |
else |
|
1045 |
I) |
|
1046 |
|> Config.put SMT_Config.infer_triggers |
|
1047 |
(Config.get ctxt smt_triggers) |
|
|
52034
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
1048 |
|> repair_monomorph_context max_mono_iters default_max_mono_iters |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
1049 |
max_new_mono_instances default_max_new_mono_instances |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
1050 |
val state = Proof.map_context (repair_context) state |
|
11b48e7a4e7e
correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents:
52031
diff
changeset
|
1051 |
val ctxt = Proof.context_of state |
| 45706 | 1052 |
val max_slices = if slice then Config.get ctxt smt_max_slices else 1 |
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1053 |
fun do_slice timeout slice outcome0 time_so_far |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1054 |
(weighted_factss as (fact_filter, weighted_facts) :: _) = |
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1055 |
let |
|
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1056 |
val timer = Timer.startRealTimer () |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
1057 |
val slice_timeout = |
| 50557 | 1058 |
if slice < max_slices andalso timeout <> NONE then |
1059 |
let val ms = timeout |> the |> Time.toMilliseconds in |
|
1060 |
Int.min (ms, |
|
1061 |
Int.max (1000 * Config.get ctxt smt_slice_min_secs, |
|
1062 |
Real.ceil (Config.get ctxt smt_slice_time_frac |
|
1063 |
* Real.fromInt ms))) |
|
1064 |
|> Time.fromMilliseconds |> SOME |
|
1065 |
end |
|
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1066 |
else |
|
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1067 |
timeout |
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1068 |
val num_facts = length weighted_facts |
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1069 |
val _ = |
|
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1070 |
if debug then |
|
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1071 |
quote name ^ " slice " ^ string_of_int slice ^ " with " ^ |
| 50557 | 1072 |
string_of_int num_facts ^ " fact" ^ plural_s num_facts ^ |
1073 |
(case slice_timeout of |
|
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1074 |
SOME timeout => " for " ^ string_of_time timeout |
| 50557 | 1075 |
| NONE => "") ^ "..." |
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1076 |
|> Output.urgent_message |
|
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1077 |
else |
|
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1078 |
() |
| 41168 | 1079 |
val birth = Timer.checkRealTimer timer |
| 41171 | 1080 |
val _ = |
|
41211
1e2e16bc0077
no need to do a super-duper atomization if Metis fails afterwards anyway
blanchet
parents:
41209
diff
changeset
|
1081 |
if debug then Output.urgent_message "Invoking SMT solver..." else () |
| 41209 | 1082 |
val (outcome, used_facts) = |
|
51204
20f6d400cc68
don't pass chained facts directly to SMT solvers -- this breaks various invariants and is never necessary
blanchet
parents:
51200
diff
changeset
|
1083 |
SMT_Solver.smt_filter_preprocess ctxt [] goal weighted_facts i |
| 50557 | 1084 |
|> SMT_Solver.smt_filter_apply (slice_timeout |> the_default one_day) |
| 41239 | 1085 |
|> (fn {outcome, used_facts} => (outcome, used_facts))
|
| 41209 | 1086 |
handle exn => if Exn.is_interrupt exn then |
1087 |
reraise exn |
|
1088 |
else |
|
|
42061
71077681eaf6
let SMT errors through -- the main reason for keeping them quiet was that the SMT bridge used to suffer from internal bugs, but these have been fixed for some time now
blanchet
parents:
42060
diff
changeset
|
1089 |
(ML_Compiler.exn_message exn |
| 41209 | 1090 |
|> SMT_Failure.Other_Failure |> SOME, []) |
| 41168 | 1091 |
val death = Timer.checkRealTimer timer |
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1092 |
val outcome0 = if is_none outcome0 then SOME outcome else outcome0 |
| 41168 | 1093 |
val time_so_far = Time.+ (time_so_far, Time.- (death, birth)) |
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1094 |
val too_many_facts_perhaps = |
|
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1095 |
case outcome of |
|
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1096 |
NONE => false |
|
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1097 |
| SOME (SMT_Failure.Counterexample _) => false |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
1098 |
| SOME SMT_Failure.Time_Out => slice_timeout <> timeout |
|
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1099 |
| SOME (SMT_Failure.Abnormal_Termination _) => true (* kind of *) |
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1100 |
| SOME SMT_Failure.Out_Of_Memory => true |
|
41211
1e2e16bc0077
no need to do a super-duper atomization if Metis fails afterwards anyway
blanchet
parents:
41209
diff
changeset
|
1101 |
| SOME (SMT_Failure.Other_Failure _) => true |
| 50557 | 1102 |
val timeout = |
1103 |
Option.map |
|
1104 |
(fn timeout => Time.- (timeout, Timer.checkRealTimer timer)) |
|
1105 |
timeout |
|
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1106 |
in |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
1107 |
if too_many_facts_perhaps andalso slice < max_slices andalso |
| 50557 | 1108 |
num_facts > 0 andalso |
1109 |
(timeout = NONE orelse Time.> (the timeout, Time.zeroTime)) then |
|
| 41169 | 1110 |
let |
|
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1111 |
val new_num_facts = |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
1112 |
Real.ceil (Config.get ctxt smt_slice_fact_frac |
|
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42642
diff
changeset
|
1113 |
* Real.fromInt num_facts) |
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1114 |
val weighted_factss as (new_fact_filter, _) :: _ = |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1115 |
weighted_factss |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1116 |
|> rotate_one |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1117 |
|> app_hd (apsnd (take new_num_facts)) |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1118 |
val show_filter = fact_filter <> new_fact_filter |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1119 |
fun num_of_facts fact_filter num_facts = |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1120 |
string_of_int num_facts ^ |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1121 |
(if show_filter then " " ^ quote fact_filter else "") ^ |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1122 |
" fact" ^ plural_s num_facts |
|
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1123 |
val _ = |
|
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1124 |
if verbose andalso is_some outcome then |
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1125 |
quote name ^ " invoked with " ^ |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1126 |
num_of_facts fact_filter num_facts ^ ": " ^ |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1127 |
string_of_failure (failure_of_smt_failure (the outcome)) ^ |
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1128 |
" Retrying with " ^ num_of_facts new_fact_filter new_num_facts ^ |
|
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1129 |
"..." |
|
42614
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1130 |
|> Output.urgent_message |
|
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1131 |
else |
|
81953e554197
make "debug" more verbose and "verbose" less verbose
blanchet
parents:
42613
diff
changeset
|
1132 |
() |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
1133 |
in |
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1134 |
do_slice timeout (slice + 1) outcome0 time_so_far weighted_factss |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
1135 |
end |
|
40553
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1136 |
else |
|
1264c9172338
pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents:
40471
diff
changeset
|
1137 |
{outcome = if is_none outcome then NONE else the outcome0,
|
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1138 |
used_facts = used_facts, used_from = map (apsnd snd) weighted_facts, |
|
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1139 |
run_time = time_so_far} |
|
40409
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents:
40370
diff
changeset
|
1140 |
end |
|
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42361
diff
changeset
|
1141 |
in do_slice timeout 1 NONE Time.zeroTime end |
|
40409
3642dc3b72e8
invoke SMT solver in a loop, with fewer and fewer facts, in case of error
blanchet
parents:
40370
diff
changeset
|
1142 |
|
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1143 |
fun run_smt_solver mode name (params as {debug, verbose, preplay_timeout, ...})
|
|
43011
5f8d74d3b297
added syntax for specifying Metis timeout (currently used only by SMT solvers)
blanchet
parents:
43006
diff
changeset
|
1144 |
minimize_command |
| 51014 | 1145 |
({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
|
|
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
1146 |
let |
|
41242
8edeb1dbbc76
run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents:
41241
diff
changeset
|
1147 |
val ctxt = Proof.context_of state |
| 51014 | 1148 |
fun weight_facts facts = |
1149 |
let val num_facts = length facts in |
|
1150 |
facts ~~ (0 upto num_facts - 1) |
|
1151 |
|> map (weight_smt_fact ctxt num_facts) |
|
1152 |
end |
|
|
51024
98fb341d32e3
distinguish MeSh and smart -- with smart, allow combinations of MaSh, MeSh, and MePo in different slices -- and use MaSh also with SMT solvers, based on evaluation
blanchet
parents:
51014
diff
changeset
|
1153 |
val weighted_factss = factss |> map (apsnd weight_facts) |
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1154 |
val {outcome, used_facts = used_pairs, used_from, run_time} =
|
| 51014 | 1155 |
smt_filter_loop name params state goal subgoal weighted_factss |
|
45781
fc2c368b5f54
use same order of facts for preplay as for actual reconstruction -- Metis sometimes exhibits very different timings depending on the order of the facts
blanchet
parents:
45707
diff
changeset
|
1156 |
val used_facts = used_pairs |> map fst |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1157 |
val outcome = outcome |> Option.map failure_of_smt_failure |
|
43261
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
1158 |
val (preplay, message, message_tail) = |
|
40184
91b4b73dbafb
proper error handling for SMT solvers in Sledgehammer
blanchet
parents:
40181
diff
changeset
|
1159 |
case outcome of |
|
91b4b73dbafb
proper error handling for SMT solvers in Sledgehammer
blanchet
parents:
40181
diff
changeset
|
1160 |
NONE => |
|
50669
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
1161 |
(Lazy.lazy (fn () => |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
1162 |
play_one_line_proof mode debug verbose preplay_timeout used_pairs |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
1163 |
state subgoal SMT |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
1164 |
(bunch_of_reconstructors false (fn desperate => |
|
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
1165 |
if desperate then liftingN else metis_default_lam_trans))), |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1166 |
fn preplay => |
|
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1167 |
let |
|
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1168 |
val one_line_params = |
|
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1169 |
(preplay, proof_banner mode name, used_facts, |
|
51200
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
1170 |
choose_minimize_command ctxt params minimize_command name |
|
260cb10aac4b
minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents:
51190
diff
changeset
|
1171 |
preplay, |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1172 |
subgoal, subgoal_count) |
| 48799 | 1173 |
val num_chained = length (#facts (Proof.goal state)) |
1174 |
in one_line_proof_text num_chained one_line_params end, |
|
|
43261
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
1175 |
if verbose then |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51998
diff
changeset
|
1176 |
"\nSMT solver real CPU time: " ^ string_of_time run_time ^ "." |
|
43261
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
1177 |
else |
|
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
1178 |
"") |
| 43166 | 1179 |
| SOME failure => |
|
50669
84c7cf36b2e0
use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents:
50667
diff
changeset
|
1180 |
(Lazy.value (Failed_to_Play plain_metis), |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
1181 |
fn _ => string_of_failure failure, "") |
|
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1182 |
in |
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1183 |
{outcome = outcome, used_facts = used_facts, used_from = used_from,
|
|
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1184 |
run_time = run_time, preplay = preplay, message = message, |
|
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1185 |
message_tail = message_tail} |
|
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1186 |
end |
|
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1187 |
|
| 45520 | 1188 |
fun run_reconstructor mode name |
|
45561
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1189 |
(params as {debug, verbose, timeout, type_enc, lam_trans, ...})
|
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
1190 |
minimize_command |
| 51010 | 1191 |
({state, subgoal, subgoal_count, factss = (_, facts) :: _, ...}
|
|
51007
4f694d52bf62
thread fact triple (MeSh, MePo, MaSh) to allow different filters in different slices
blanchet
parents:
51005
diff
changeset
|
1192 |
: prover_problem) = |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1193 |
let |
| 45520 | 1194 |
val reconstr = |
1195 |
if name = metisN then |
|
1196 |
Metis (type_enc |> the_default (hd partial_type_encs), |
|
1197 |
lam_trans |> the_default metis_default_lam_trans) |
|
1198 |
else if name = smtN then |
|
1199 |
SMT |
|
1200 |
else |
|
1201 |
raise Fail ("unknown reconstructor: " ^ quote name)
|
|
|
51005
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents:
50927
diff
changeset
|
1202 |
val used_facts = facts |> map fst |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1203 |
in |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
1204 |
case play_one_line_proof (if mode = Minimize then Normal else mode) debug |
|
51005
ce4290c33d73
eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents:
50927
diff
changeset
|
1205 |
verbose timeout facts state subgoal reconstr |
| 45520 | 1206 |
[reconstr] of |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1207 |
play as Played (_, time) => |
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1208 |
{outcome = NONE, used_facts = used_facts, used_from = facts,
|
|
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1209 |
run_time = time, preplay = Lazy.value play, |
|
45561
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1210 |
message = |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1211 |
fn play => |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1212 |
let |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1213 |
val (_, override_params) = extract_reconstructor params reconstr |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1214 |
val one_line_params = |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1215 |
(play, proof_banner mode name, used_facts, |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1216 |
minimize_command override_params name, subgoal, |
|
57227eedce81
don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents:
45560
diff
changeset
|
1217 |
subgoal_count) |
| 48799 | 1218 |
val num_chained = length (#facts (Proof.goal state)) |
1219 |
in one_line_proof_text num_chained one_line_params end, |
|
|
43261
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents:
43259
diff
changeset
|
1220 |
message_tail = ""} |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1221 |
| play => |
| 43166 | 1222 |
let |
1223 |
val failure = case play of Failed_to_Play _ => GaveUp | _ => TimedOut |
|
1224 |
in |
|
|
51009
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1225 |
{outcome = SOME failure, used_facts = [], used_from = [],
|
|
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
blanchet
parents:
51007
diff
changeset
|
1226 |
run_time = Time.zeroTime, preplay = Lazy.value play, |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51879
diff
changeset
|
1227 |
message = fn _ => string_of_failure failure, message_tail = ""} |
|
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1228 |
end |
|
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1229 |
end |
|
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
43044
diff
changeset
|
1230 |
|
| 43021 | 1231 |
fun get_prover ctxt mode name = |
| 42361 | 1232 |
let val thy = Proof_Context.theory_of ctxt in |
|
45379
0147a4348ca1
try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents:
45378
diff
changeset
|
1233 |
if is_reconstructor name then run_reconstructor mode name |
|
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47531
diff
changeset
|
1234 |
else if is_atp thy name then run_atp mode name (get_atp thy name ()) |
|
43052
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1235 |
else if is_smt_prover ctxt name then run_smt_solver mode name |
|
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
blanchet
parents:
43051
diff
changeset
|
1236 |
else error ("No such prover: " ^ name ^ ".")
|
|
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40723
diff
changeset
|
1237 |
end |
|
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
1238 |
|
| 28582 | 1239 |
end; |