| author | blanchet | 
| Thu, 25 Aug 2011 00:00:36 +0200 | |
| changeset 44487 | 0989d8deab69 | 
| parent 44462 | d9a657c44380 | 
| child 44542 | 3f5fd3635281 | 
| permissions | -rw-r--r-- | 
| 32564 | 1  | 
(* Title: HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML  | 
2  | 
Author: Jasmin Blanchette and Sascha Boehme and Tobias Nipkow, TU Munich  | 
|
| 
32385
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
5  | 
structure Mirabelle_Sledgehammer : MIRABELLE_ACTION =  | 
| 
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
6  | 
struct  | 
| 
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
7  | 
|
| 32521 | 8  | 
val proverK = "prover"  | 
| 32541 | 9  | 
val prover_timeoutK = "prover_timeout"  | 
| 32521 | 10  | 
val keepK = "keep"  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43569 
diff
changeset
 | 
11  | 
val type_encK = "type_enc"  | 
| 44089 | 12  | 
val soundK = "sound"  | 
| 
42725
 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 
blanchet 
parents: 
42642 
diff
changeset
 | 
13  | 
val slicingK = "slicing"  | 
| 
43827
 
62d64709af3b
added option to control which lambda translation to use (for experiments)
 
blanchet 
parents: 
43710 
diff
changeset
 | 
14  | 
val lambda_translationK = "lambda_translation"  | 
| 
42725
 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 
blanchet 
parents: 
42642 
diff
changeset
 | 
15  | 
val e_weight_methodK = "e_weight_method"  | 
| 44099 | 16  | 
val force_sosK = "force_sos"  | 
| 41752 | 17  | 
val max_relevantK = "max_relevant"  | 
| 
44431
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
18  | 
val max_callsK = "max_calls"  | 
| 32525 | 19  | 
val minimizeK = "minimize"  | 
20  | 
val minimize_timeoutK = "minimize_timeout"  | 
|
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
21  | 
val metis_ftK = "metis_ft"  | 
| 
41357
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
22  | 
val reconstructorK = "reconstructor"  | 
| 32521 | 23  | 
|
| 
44448
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
24  | 
val preplay_timeout = "4"  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
25  | 
|
| 32521 | 26  | 
fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: "  | 
| 32525 | 27  | 
fun minimize_tag id = "#" ^ string_of_int id ^ " minimize (sledgehammer): "  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
28  | 
fun reconstructor_tag reconstructor id =  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
29  | 
"#" ^ string_of_int id ^ " " ^ (!reconstructor) ^ " (sledgehammer): "  | 
| 32521 | 30  | 
|
| 32525 | 31  | 
val separator = "-----"  | 
32  | 
||
| 32521 | 33  | 
|
| 32549 | 34  | 
datatype sh_data = ShData of {
 | 
35  | 
calls: int,  | 
|
36  | 
success: int,  | 
|
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
37  | 
nontriv_calls: int,  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
38  | 
nontriv_success: int,  | 
| 32585 | 39  | 
lemmas: int,  | 
| 32818 | 40  | 
max_lems: int,  | 
| 32549 | 41  | 
time_isa: int,  | 
| 40062 | 42  | 
time_prover: int,  | 
43  | 
time_prover_fail: int}  | 
|
| 32549 | 44  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
45  | 
datatype re_data = ReData of {
 | 
| 32549 | 46  | 
calls: int,  | 
47  | 
success: int,  | 
|
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
48  | 
nontriv_calls: int,  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
49  | 
nontriv_success: int,  | 
| 32676 | 50  | 
proofs: int,  | 
| 32549 | 51  | 
time: int,  | 
| 32550 | 52  | 
timeout: int,  | 
| 32990 | 53  | 
lemmas: int * int * int,  | 
| 
39341
 
d2b981a0429a
indicate triviality in the list of proved things
 
blanchet 
parents: 
39340 
diff
changeset
 | 
54  | 
posns: (Position.T * bool) list  | 
| 32550 | 55  | 
}  | 
| 32549 | 56  | 
|
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
57  | 
datatype min_data = MinData of {
 | 
| 32609 | 58  | 
succs: int,  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35830 
diff
changeset
 | 
59  | 
ab_ratios: int  | 
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
60  | 
}  | 
| 32521 | 61  | 
|
| 32818 | 62  | 
fun make_sh_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
63  | 
(calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa,  | 
| 40062 | 64  | 
time_prover,time_prover_fail) =  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
65  | 
  ShData{calls=calls, success=success, nontriv_calls=nontriv_calls,
 | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
66  | 
nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems,  | 
| 40062 | 67  | 
time_isa=time_isa, time_prover=time_prover,  | 
68  | 
time_prover_fail=time_prover_fail}  | 
|
| 32521 | 69  | 
|
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35830 
diff
changeset
 | 
70  | 
fun make_min_data (succs, ab_ratios) =  | 
| 
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35830 
diff
changeset
 | 
71  | 
  MinData{succs=succs, ab_ratios=ab_ratios}
 | 
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
72  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
73  | 
fun make_re_data (calls,success,nontriv_calls,nontriv_success,proofs,time,  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
74  | 
timeout,lemmas,posns) =  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
75  | 
  ReData{calls=calls, success=success, nontriv_calls=nontriv_calls,
 | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
76  | 
nontriv_success=nontriv_success, proofs=proofs, time=time,  | 
| 32990 | 77  | 
timeout=timeout, lemmas=lemmas, posns=posns}  | 
| 32549 | 78  | 
|
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
79  | 
val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0, 0, 0)  | 
| 
35871
 
c93bda4fdf15
remove the iteration counter from Sledgehammer's minimizer
 
blanchet 
parents: 
35867 
diff
changeset
 | 
80  | 
val empty_min_data = make_min_data (0, 0)  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
81  | 
val empty_re_data = make_re_data (0, 0, 0, 0, 0, 0, 0, (0,0,0), [])  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
82  | 
|
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
83  | 
fun tuple_of_sh_data (ShData {calls, success, nontriv_calls, nontriv_success,
 | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
84  | 
lemmas, max_lems, time_isa,  | 
| 40062 | 85  | 
time_prover, time_prover_fail}) = (calls, success, nontriv_calls,  | 
86  | 
nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail)  | 
|
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
87  | 
|
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35830 
diff
changeset
 | 
88  | 
fun tuple_of_min_data (MinData {succs, ab_ratios}) = (succs, ab_ratios)
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
89  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
90  | 
fun tuple_of_re_data (ReData {calls, success, nontriv_calls, nontriv_success,
 | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
91  | 
proofs, time, timeout, lemmas, posns}) = (calls, success, nontriv_calls,  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
92  | 
nontriv_success, proofs, time, timeout, lemmas, posns)  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
93  | 
|
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
94  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
95  | 
datatype reconstructor_mode =  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
96  | 
Unminimized | Minimized | UnminimizedFT | MinimizedFT  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
97  | 
|
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
98  | 
datatype data = Data of {
 | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
99  | 
sh: sh_data,  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
100  | 
min: min_data,  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
101  | 
re_u: re_data, (* reconstructor with unminimized set of lemmas *)  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
102  | 
re_m: re_data, (* reconstructor with minimized set of lemmas *)  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
103  | 
re_uft: re_data, (* reconstructor with unminimized set of lemmas and fully-typed *)  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
104  | 
re_mft: re_data, (* reconstructor with minimized set of lemmas and fully-typed *)  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
105  | 
mini: bool (* with minimization *)  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
106  | 
}  | 
| 32521 | 107  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
108  | 
fun make_data (sh, min, re_u, re_m, re_uft, re_mft, mini) =  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
109  | 
  Data {sh=sh, min=min, re_u=re_u, re_m=re_m, re_uft=re_uft, re_mft=re_mft,
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
110  | 
mini=mini}  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
111  | 
|
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
112  | 
val empty_data = make_data (empty_sh_data, empty_min_data,  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
113  | 
empty_re_data, empty_re_data, empty_re_data, empty_re_data, false)  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
114  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
115  | 
fun map_sh_data f (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
116  | 
let val sh' = make_sh_data (f (tuple_of_sh_data sh))  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
117  | 
in make_data (sh', min, re_u, re_m, re_uft, re_mft, mini) end  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
118  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
119  | 
fun map_min_data f (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
120  | 
let val min' = make_min_data (f (tuple_of_min_data min))  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
121  | 
in make_data (sh, min', re_u, re_m, re_uft, re_mft, mini) end  | 
| 32521 | 122  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
123  | 
fun map_re_data f m (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
124  | 
let  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
125  | 
fun map_me g Unminimized (u, m, uft, mft) = (g u, m, uft, mft)  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
126  | 
| map_me g Minimized (u, m, uft, mft) = (u, g m, uft, mft)  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
127  | 
| map_me g UnminimizedFT (u, m, uft, mft) = (u, m, g uft, mft)  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
128  | 
| map_me g MinimizedFT (u, m, uft, mft) = (u, m, uft, g mft)  | 
| 32521 | 129  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
130  | 
val f' = make_re_data o f o tuple_of_re_data  | 
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
131  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
132  | 
val (re_u', re_m', re_uft', re_mft') =  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
133  | 
map_me f' m (re_u, re_m, re_uft, re_mft)  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
134  | 
in make_data (sh, min, re_u', re_m', re_uft', re_mft', mini) end  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
135  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
136  | 
fun set_mini mini (Data {sh, min, re_u, re_m, re_uft, re_mft, ...}) =
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
137  | 
make_data (sh, min, re_u, re_m, re_uft, re_mft, mini)  | 
| 32990 | 138  | 
|
139  | 
fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n));  | 
|
| 32521 | 140  | 
|
| 32818 | 141  | 
val inc_sh_calls = map_sh_data  | 
| 40062 | 142  | 
(fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)  | 
143  | 
=> (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))  | 
|
| 32549 | 144  | 
|
| 32818 | 145  | 
val inc_sh_success = map_sh_data  | 
| 40062 | 146  | 
(fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)  | 
147  | 
=> (calls, success + 1, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail))  | 
|
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
148  | 
|
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
149  | 
val inc_sh_nontriv_calls = map_sh_data  | 
| 40062 | 150  | 
(fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)  | 
151  | 
=> (calls, success, nontriv_calls + 1, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))  | 
|
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
152  | 
|
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
153  | 
val inc_sh_nontriv_success = map_sh_data  | 
| 40062 | 154  | 
(fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)  | 
155  | 
=> (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_prover, time_prover_fail))  | 
|
| 32585 | 156  | 
|
| 32818 | 157  | 
fun inc_sh_lemmas n = map_sh_data  | 
| 40062 | 158  | 
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)  | 
159  | 
=> (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_prover,time_prover_fail))  | 
|
| 32521 | 160  | 
|
| 32818 | 161  | 
fun inc_sh_max_lems n = map_sh_data  | 
| 40062 | 162  | 
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)  | 
163  | 
=> (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_prover,time_prover_fail))  | 
|
| 32549 | 164  | 
|
| 32818 | 165  | 
fun inc_sh_time_isa t = map_sh_data  | 
| 40062 | 166  | 
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)  | 
167  | 
=> (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_prover,time_prover_fail))  | 
|
| 32818 | 168  | 
|
| 40062 | 169  | 
fun inc_sh_time_prover t = map_sh_data  | 
170  | 
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)  | 
|
171  | 
=> (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover + t,time_prover_fail))  | 
|
| 32521 | 172  | 
|
| 40062 | 173  | 
fun inc_sh_time_prover_fail t = map_sh_data  | 
174  | 
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)  | 
|
175  | 
=> (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail + t))  | 
|
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
176  | 
|
| 32818 | 177  | 
val inc_min_succs = map_min_data  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35830 
diff
changeset
 | 
178  | 
(fn (succs,ab_ratios) => (succs+1, ab_ratios))  | 
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
179  | 
|
| 32818 | 180  | 
fun inc_min_ab_ratios r = map_min_data  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35830 
diff
changeset
 | 
181  | 
(fn (succs, ab_ratios) => (succs, ab_ratios+r))  | 
| 32549 | 182  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
183  | 
val inc_reconstructor_calls = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
184  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
185  | 
=> (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))  | 
| 32533 | 186  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
187  | 
val inc_reconstructor_success = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
188  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
189  | 
=> (calls, success + 1, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
190  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
191  | 
val inc_reconstructor_nontriv_calls = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
192  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
193  | 
=> (calls, success, nontriv_calls + 1, nontriv_success, proofs, time, timeout, lemmas,posns))  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
194  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
195  | 
val inc_reconstructor_nontriv_success = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
196  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
197  | 
=> (calls, success, nontriv_calls, nontriv_success + 1, proofs, time, timeout, lemmas,posns))  | 
| 32676 | 198  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
199  | 
val inc_reconstructor_proofs = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
200  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
201  | 
=> (calls, success, nontriv_calls, nontriv_success, proofs + 1, time, timeout, lemmas,posns))  | 
| 
32551
 
421323205efd
position information is now passed to all actions;
 
nipkow 
parents: 
32550 
diff
changeset
 | 
202  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
203  | 
fun inc_reconstructor_time m t = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
204  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
205  | 
=> (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns)) m  | 
| 32536 | 206  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
207  | 
val inc_reconstructor_timeout = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
208  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
209  | 
=> (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout + 1, lemmas,posns))  | 
| 32549 | 210  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
211  | 
fun inc_reconstructor_lemmas m n = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
212  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
213  | 
=> (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, inc_max n lemmas, posns)) m  | 
| 
32551
 
421323205efd
position information is now passed to all actions;
 
nipkow 
parents: 
32550 
diff
changeset
 | 
214  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
215  | 
fun inc_reconstructor_posns m pos = map_re_data  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
216  | 
(fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)  | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
217  | 
=> (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns)) m  | 
| 32521 | 218  | 
|
| 44090 | 219  | 
val str0 = string_of_int o the_default 0  | 
220  | 
||
| 32521 | 221  | 
local  | 
222  | 
||
223  | 
val str = string_of_int  | 
|
224  | 
val str3 = Real.fmt (StringCvt.FIX (SOME 3))  | 
|
225  | 
fun percentage a b = string_of_int (a * 100 div b)  | 
|
226  | 
fun time t = Real.fromInt t / 1000.0  | 
|
227  | 
fun avg_time t n =  | 
|
228  | 
if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0  | 
|
229  | 
||
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
230  | 
fun log_sh_data log  | 
| 40062 | 231  | 
(calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail) =  | 
| 32818 | 232  | 
 (log ("Total number of sledgehammer calls: " ^ str calls);
 | 
233  | 
  log ("Number of successful sledgehammer calls: " ^ str success);
 | 
|
234  | 
  log ("Number of sledgehammer lemmas: " ^ str lemmas);
 | 
|
235  | 
  log ("Max number of sledgehammer lemmas: " ^ str max_lems);
 | 
|
236  | 
  log ("Success rate: " ^ percentage success calls ^ "%");
 | 
|
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
237  | 
  log ("Total number of nontrivial sledgehammer calls: " ^ str nontriv_calls);
 | 
| 
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
238  | 
  log ("Number of successful nontrivial sledgehammer calls: " ^ str nontriv_success);
 | 
| 32818 | 239  | 
  log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
 | 
| 40062 | 240  | 
  log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_prover));
 | 
241  | 
  log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_prover_fail));
 | 
|
| 32536 | 242  | 
  log ("Average time for sledgehammer calls (Isabelle): " ^
 | 
| 32818 | 243  | 
str3 (avg_time time_isa calls));  | 
| 32533 | 244  | 
  log ("Average time for successful sledgehammer calls (ATP): " ^
 | 
| 40062 | 245  | 
str3 (avg_time time_prover success));  | 
| 32536 | 246  | 
  log ("Average time for failed sledgehammer calls (ATP): " ^
 | 
| 40062 | 247  | 
str3 (avg_time time_prover_fail (calls - success)))  | 
| 32533 | 248  | 
)  | 
| 32521 | 249  | 
|
| 
39341
 
d2b981a0429a
indicate triviality in the list of proved things
 
blanchet 
parents: 
39340 
diff
changeset
 | 
250  | 
fun str_of_pos (pos, triv) =  | 
| 44090 | 251  | 
str0 (Position.line_of pos) (* ^ ":" ^ str0 (Position.offset_of pos) *) ^  | 
252  | 
(if triv then "[T]" else "")  | 
|
| 
32551
 
421323205efd
position information is now passed to all actions;
 
nipkow 
parents: 
32550 
diff
changeset
 | 
253  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
254  | 
fun log_re_data log tag sh_calls (re_calls, re_success, re_nontriv_calls,  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
255  | 
re_nontriv_success, re_proofs, re_time, re_timeout,  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
256  | 
(lemmas, lems_sos, lems_max), re_posns) =  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
257  | 
 (log ("Total number of " ^ tag ^ "reconstructor calls: " ^ str re_calls);
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
258  | 
  log ("Number of successful " ^ tag ^ "reconstructor calls: " ^ str re_success ^
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
259  | 
" (proof: " ^ str re_proofs ^ ")");  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
260  | 
  log ("Number of " ^ tag ^ "reconstructor timeouts: " ^ str re_timeout);
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
261  | 
  log ("Success rate: " ^ percentage re_success sh_calls ^ "%");
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
262  | 
  log ("Total number of nontrivial " ^ tag ^ "reconstructor calls: " ^ str re_nontriv_calls);
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
263  | 
  log ("Number of successful nontrivial " ^ tag ^ "reconstructor calls: " ^ str re_nontriv_success ^
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
264  | 
" (proof: " ^ str re_proofs ^ ")");  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
265  | 
  log ("Number of successful " ^ tag ^ "reconstructor lemmas: " ^ str lemmas);
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
266  | 
  log ("SOS of successful " ^ tag ^ "reconstructor lemmas: " ^ str lems_sos);
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
267  | 
  log ("Max number of successful " ^ tag ^ "reconstructor lemmas: " ^ str lems_max);
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
268  | 
  log ("Total time for successful " ^ tag ^ "reconstructor calls: " ^ str3 (time re_time));
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
269  | 
  log ("Average time for successful " ^ tag ^ "reconstructor calls: " ^
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
270  | 
str3 (avg_time re_time re_success));  | 
| 
32551
 
421323205efd
position information is now passed to all actions;
 
nipkow 
parents: 
32550 
diff
changeset
 | 
271  | 
if tag=""  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
272  | 
  then log ("Proved: " ^ space_implode " " (map str_of_pos re_posns))
 | 
| 
32551
 
421323205efd
position information is now passed to all actions;
 
nipkow 
parents: 
32550 
diff
changeset
 | 
273  | 
else ()  | 
| 
 
421323205efd
position information is now passed to all actions;
 
nipkow 
parents: 
32550 
diff
changeset
 | 
274  | 
)  | 
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
275  | 
|
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35830 
diff
changeset
 | 
276  | 
fun log_min_data log (succs, ab_ratios) =  | 
| 32609 | 277  | 
  (log ("Number of successful minimizations: " ^ string_of_int succs);
 | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents: 
35830 
diff
changeset
 | 
278  | 
   log ("After/before ratios: " ^ string_of_int ab_ratios)
 | 
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
279  | 
)  | 
| 
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
280  | 
|
| 32521 | 281  | 
in  | 
282  | 
||
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
283  | 
fun log_data id log (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
284  | 
let  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
285  | 
    val ShData {calls=sh_calls, ...} = sh
 | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
286  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
287  | 
    fun app_if (ReData {calls, ...}) f = if calls > 0 then f () else ()
 | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
288  | 
fun log_re tag m =  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
289  | 
log_re_data log tag sh_calls (tuple_of_re_data m)  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
290  | 
fun log_reconstructor (tag1, m1) (tag2, m2) = app_if m1 (fn () =>  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
291  | 
(log_re tag1 m1; log ""; app_if m2 (fn () => log_re tag2 m2)))  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
292  | 
in  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
293  | 
if sh_calls > 0  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
294  | 
then  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
295  | 
     (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
 | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
296  | 
log_sh_data log (tuple_of_sh_data sh);  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
297  | 
log "";  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
298  | 
if not mini  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
299  | 
      then log_reconstructor ("", re_u) ("fully-typed ", re_uft)
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
300  | 
else  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
301  | 
app_if re_u (fn () =>  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
302  | 
         (log_reconstructor ("unminimized ", re_u) ("unminimized fully-typed ", re_uft);
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
303  | 
log "";  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
304  | 
app_if re_m (fn () =>  | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
305  | 
(log_min_data log (tuple_of_min_data min); log "";  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
306  | 
             log_reconstructor ("", re_m) ("fully-typed ", re_mft))))))
 | 
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
307  | 
else ()  | 
| 
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
308  | 
end  | 
| 32521 | 309  | 
|
310  | 
end  | 
|
311  | 
||
312  | 
||
313  | 
(* Warning: we implicitly assume single-threaded execution here! *)  | 
|
| 32740 | 314  | 
val data = Unsynchronized.ref ([] : (int * data) list)  | 
| 32521 | 315  | 
|
| 32740 | 316  | 
fun init id thy = (Unsynchronized.change data (cons (id, empty_data)); thy)  | 
| 
32567
 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 
wenzelm 
parents: 
32564 
diff
changeset
 | 
317  | 
fun done id ({log, ...}: Mirabelle.done_args) =
 | 
| 32521 | 318  | 
AList.lookup (op =) (!data) id  | 
319  | 
|> Option.map (log_data id log)  | 
|
320  | 
|> K ()  | 
|
321  | 
||
| 32740 | 322  | 
fun change_data id f = (Unsynchronized.change data (AList.map_entry (op =) id f); ())  | 
| 32521 | 323  | 
|
324  | 
||
| 
42444
 
8e5438dc70bb
cleanup: get rid of "may_slice" arguments without changing semantics
 
blanchet 
parents: 
42443 
diff
changeset
 | 
325  | 
fun get_prover ctxt args =  | 
| 
33016
 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 
boehmes 
parents: 
32991 
diff
changeset
 | 
326  | 
let  | 
| 40062 | 327  | 
fun default_prover_name () =  | 
| 40069 | 328  | 
hd (#provers (Sledgehammer_Isar.default_params ctxt []))  | 
| 
33016
 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 
boehmes 
parents: 
32991 
diff
changeset
 | 
329  | 
handle Empty => error "No ATP available."  | 
| 
41087
 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 
blanchet 
parents: 
40983 
diff
changeset
 | 
330  | 
fun get_prover name =  | 
| 43021 | 331  | 
(name, Sledgehammer_Run.get_minimizing_prover ctxt  | 
332  | 
Sledgehammer_Provers.Normal name)  | 
|
| 
33016
 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 
boehmes 
parents: 
32991 
diff
changeset
 | 
333  | 
in  | 
| 
 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 
boehmes 
parents: 
32991 
diff
changeset
 | 
334  | 
(case AList.lookup (op =) args proverK of  | 
| 40062 | 335  | 
SOME name => get_prover name  | 
336  | 
| NONE => get_prover (default_prover_name ()))  | 
|
| 
33016
 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 
boehmes 
parents: 
32991 
diff
changeset
 | 
337  | 
end  | 
| 32525 | 338  | 
|
| 43088 | 339  | 
type locality = ATP_Translate.locality  | 
| 
38752
 
6628adcae4a7
consider "locality" when assigning weights to facts
 
blanchet 
parents: 
38700 
diff
changeset
 | 
340  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
341  | 
(* hack *)  | 
| 
41357
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
342  | 
fun reconstructor_from_msg args msg =  | 
| 
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
343  | 
(case AList.lookup (op =) args reconstructorK of  | 
| 
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
344  | 
SOME name => name  | 
| 
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
345  | 
| NONE =>  | 
| 
43228
 
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
 
blanchet 
parents: 
43212 
diff
changeset
 | 
346  | 
if String.isSubstring "metis (full_types)" msg then "metis (full_types)"  | 
| 
 
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
 
blanchet 
parents: 
43212 
diff
changeset
 | 
347  | 
else if String.isSubstring "metis (no_types)" msg then "metis (no_types)"  | 
| 
41357
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
348  | 
else if String.isSubstring "metis" msg then "metis"  | 
| 
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
349  | 
else "smt")  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
350  | 
|
| 32521 | 351  | 
local  | 
352  | 
||
| 32536 | 353  | 
datatype sh_result =  | 
| 
38752
 
6628adcae4a7
consider "locality" when assigning weights to facts
 
blanchet 
parents: 
38700 
diff
changeset
 | 
354  | 
SH_OK of int * int * (string * locality) list |  | 
| 32536 | 355  | 
SH_FAIL of int * int |  | 
356  | 
SH_ERROR  | 
|
357  | 
||
| 44089 | 358  | 
fun run_sh prover_name prover type_enc sound max_relevant slicing  | 
| 44099 | 359  | 
lambda_translation e_weight_method force_sos hard_timeout timeout dir  | 
360  | 
pos st =  | 
|
| 32521 | 361  | 
let  | 
| 38998 | 362  | 
    val {context = ctxt, facts = chained_ths, goal} = Proof.goal st
 | 
363  | 
val i = 1  | 
|
| 44090 | 364  | 
fun set_file_name (SOME dir) =  | 
| 
41337
 
263fe1670067
mechanism to keep SMT input and output files around in Mirabelle
 
blanchet 
parents: 
41276 
diff
changeset
 | 
365  | 
Config.put Sledgehammer_Provers.dest_dir dir  | 
| 44090 | 366  | 
#> Config.put Sledgehammer_Provers.problem_prefix  | 
| 44424 | 367  | 
          ("prob_" ^ str0 (Position.line_of pos) ^ "__")
 | 
| 
41337
 
263fe1670067
mechanism to keep SMT input and output files around in Mirabelle
 
blanchet 
parents: 
41276 
diff
changeset
 | 
368  | 
#> Config.put SMT_Config.debug_files  | 
| 43088 | 369  | 
(dir ^ "/" ^ Name.desymbolize false (ATP_Util.timestamp ()) ^ "_"  | 
| 41338 | 370  | 
^ serial_string ())  | 
| 44090 | 371  | 
| set_file_name NONE = I  | 
| 39321 | 372  | 
val st' =  | 
373  | 
st |> Proof.map_context  | 
|
| 44090 | 374  | 
(set_file_name dir  | 
| 
43828
 
e07a2c4cbad8
move lambda translation option from ATP to Sledgehammer, to avoid accidentally breaking Metis (its reconstruction code can only deal with combinators)
 
blanchet 
parents: 
43827 
diff
changeset
 | 
375  | 
#> (Option.map (Config.put  | 
| 
 
e07a2c4cbad8
move lambda translation option from ATP to Sledgehammer, to avoid accidentally breaking Metis (its reconstruction code can only deal with combinators)
 
blanchet 
parents: 
43827 
diff
changeset
 | 
376  | 
Sledgehammer_Provers.atp_lambda_translation)  | 
| 
43827
 
62d64709af3b
added option to control which lambda translation to use (for experiments)
 
blanchet 
parents: 
43710 
diff
changeset
 | 
377  | 
lambda_translation |> the_default I)  | 
| 
42725
 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 
blanchet 
parents: 
42642 
diff
changeset
 | 
378  | 
#> (Option.map (Config.put ATP_Systems.e_weight_method)  | 
| 
 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 
blanchet 
parents: 
42642 
diff
changeset
 | 
379  | 
e_weight_method |> the_default I)  | 
| 44099 | 380  | 
#> (Option.map (Config.put ATP_Systems.force_sos)  | 
381  | 
force_sos |> the_default I)  | 
|
| 
44431
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
382  | 
#> Config.put Sledgehammer_Provers.measure_run_time true  | 
| 
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
383  | 
#> Config.put Sledgehammer_Provers.atp_sound_modulo_infiniteness false)  | 
| 42642 | 384  | 
    val params as {relevance_thresholds, max_relevant, slicing, ...} =
 | 
| 40069 | 385  | 
Sledgehammer_Isar.default_params ctxt  | 
| 
40554
 
ff446d5e9a62
turn on Sledgehammer verbosity so we can track down crashes
 
blanchet 
parents: 
40526 
diff
changeset
 | 
386  | 
          [("verbose", "true"),
 | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43569 
diff
changeset
 | 
387  | 
           ("type_enc", type_enc),
 | 
| 44089 | 388  | 
           ("sound", sound),
 | 
| 
44448
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
389  | 
           ("preplay_timeout", preplay_timeout),
 | 
| 41752 | 390  | 
           ("max_relevant", max_relevant),
 | 
| 
42725
 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 
blanchet 
parents: 
42642 
diff
changeset
 | 
391  | 
           ("slicing", slicing),
 | 
| 
44448
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
392  | 
           ("timeout", string_of_int timeout),
 | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
393  | 
           ("preplay_timeout", preplay_timeout)]
 | 
| 40062 | 394  | 
val default_max_relevant =  | 
| 
42443
 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 
blanchet 
parents: 
42361 
diff
changeset
 | 
395  | 
Sledgehammer_Provers.default_max_relevant_for_prover ctxt slicing  | 
| 
 
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
 
blanchet 
parents: 
42361 
diff
changeset
 | 
396  | 
prover_name  | 
| 
42952
 
96f62b77748f
tuning -- the "appropriate" terminology is inspired from TPTP
 
blanchet 
parents: 
42944 
diff
changeset
 | 
397  | 
val is_appropriate_prop =  | 
| 
 
96f62b77748f
tuning -- the "appropriate" terminology is inspired from TPTP
 
blanchet 
parents: 
42944 
diff
changeset
 | 
398  | 
Sledgehammer_Provers.is_appropriate_prop_for_prover ctxt prover_name  | 
| 
40369
 
53dca3bd4250
use the SMT integration's official list of built-ins
 
blanchet 
parents: 
40301 
diff
changeset
 | 
399  | 
val is_built_in_const =  | 
| 
41087
 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 
blanchet 
parents: 
40983 
diff
changeset
 | 
400  | 
Sledgehammer_Provers.is_built_in_const_for_prover ctxt prover_name  | 
| 
40941
 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 
blanchet 
parents: 
40694 
diff
changeset
 | 
401  | 
val relevance_fudge =  | 
| 
41087
 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 
blanchet 
parents: 
40983 
diff
changeset
 | 
402  | 
Sledgehammer_Provers.relevance_fudge_for_prover ctxt prover_name  | 
| 
40070
 
bdb890782d4a
replaced references with proper record that's threaded through
 
blanchet 
parents: 
40069 
diff
changeset
 | 
403  | 
    val relevance_override = {add = [], del = [], only = false}
 | 
| 43088 | 404  | 
val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal i  | 
| 
32574
 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 
boehmes 
parents: 
32571 
diff
changeset
 | 
405  | 
val time_limit =  | 
| 
 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 
boehmes 
parents: 
32571 
diff
changeset
 | 
406  | 
(case hard_timeout of  | 
| 
 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 
boehmes 
parents: 
32571 
diff
changeset
 | 
407  | 
NONE => I  | 
| 
 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 
boehmes 
parents: 
32571 
diff
changeset
 | 
408  | 
| SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))  | 
| 
42953
 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 
blanchet 
parents: 
42952 
diff
changeset
 | 
409  | 
fun failed failure =  | 
| 
43052
 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 
blanchet 
parents: 
43051 
diff
changeset
 | 
410  | 
      ({outcome = SOME failure, used_facts = [], run_time_in_msecs = NONE,
 | 
| 43166 | 411  | 
preplay = K (ATP_Reconstruct.Failed_to_Play ATP_Reconstruct.Metis),  | 
| 
43261
 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 
blanchet 
parents: 
43228 
diff
changeset
 | 
412  | 
message = K "", message_tail = ""}, ~1)  | 
| 
 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 
blanchet 
parents: 
43228 
diff
changeset
 | 
413  | 
    val ({outcome, used_facts, run_time_in_msecs, preplay, message,
 | 
| 
 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 
blanchet 
parents: 
43228 
diff
changeset
 | 
414  | 
message_tail} : Sledgehammer_Provers.prover_result,  | 
| 41275 | 415  | 
time_isa) = time_limit (Mirabelle.cpu_time (fn () =>  | 
416  | 
let  | 
|
| 
42953
 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 
blanchet 
parents: 
42952 
diff
changeset
 | 
417  | 
val _ = if is_appropriate_prop concl_t then ()  | 
| 
 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 
blanchet 
parents: 
42952 
diff
changeset
 | 
418  | 
else raise Fail "inappropriate"  | 
| 41275 | 419  | 
val facts =  | 
| 
43351
 
b19d95b4d736
compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
 
blanchet 
parents: 
43261 
diff
changeset
 | 
420  | 
Sledgehammer_Filter.nearly_all_facts ctxt relevance_override  | 
| 
 
b19d95b4d736
compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
 
blanchet 
parents: 
43261 
diff
changeset
 | 
421  | 
chained_ths hyp_ts concl_t  | 
| 
 
b19d95b4d736
compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
 
blanchet 
parents: 
43261 
diff
changeset
 | 
422  | 
|> filter (is_appropriate_prop o prop_of o snd)  | 
| 
 
b19d95b4d736
compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
 
blanchet 
parents: 
43261 
diff
changeset
 | 
423  | 
|> Sledgehammer_Filter.relevant_facts ctxt relevance_thresholds  | 
| 
 
b19d95b4d736
compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
 
blanchet 
parents: 
43261 
diff
changeset
 | 
424  | 
(the_default default_max_relevant max_relevant)  | 
| 
 
b19d95b4d736
compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
 
blanchet 
parents: 
43261 
diff
changeset
 | 
425  | 
is_built_in_const relevance_fudge relevance_override  | 
| 
 
b19d95b4d736
compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
 
blanchet 
parents: 
43261 
diff
changeset
 | 
426  | 
chained_ths hyp_ts concl_t  | 
| 41275 | 427  | 
val problem =  | 
428  | 
          {state = st', goal = goal, subgoal = i,
 | 
|
429  | 
subgoal_count = Sledgehammer_Util.subgoal_count st,  | 
|
430  | 
facts = facts |> map Sledgehammer_Provers.Untranslated_Fact,  | 
|
| 41741 | 431  | 
smt_filter = NONE}  | 
| 43051 | 432  | 
in prover params (K (K "")) problem end)) ()  | 
| 
42953
 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 
blanchet 
parents: 
42952 
diff
changeset
 | 
433  | 
handle TimeLimit.TimeOut => failed ATP_Proof.TimedOut  | 
| 
 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 
blanchet 
parents: 
42952 
diff
changeset
 | 
434  | 
| Fail "inappropriate" => failed ATP_Proof.Inappropriate  | 
| 
40374
 
443b426e05ea
make Mirabelle work correctly if the prover (e.g. the SMT solver) returns no timing information
 
blanchet 
parents: 
40372 
diff
changeset
 | 
435  | 
val time_prover = run_time_in_msecs |> the_default ~1  | 
| 
43261
 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 
blanchet 
parents: 
43228 
diff
changeset
 | 
436  | 
val msg = message (preplay ()) ^ message_tail  | 
| 32521 | 437  | 
in  | 
| 36405 | 438  | 
case outcome of  | 
| 
43052
 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 
blanchet 
parents: 
43051 
diff
changeset
 | 
439  | 
NONE => (msg, SH_OK (time_isa, time_prover, used_facts))  | 
| 
 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 
blanchet 
parents: 
43051 
diff
changeset
 | 
440  | 
| SOME _ => (msg, SH_FAIL (time_isa, time_prover))  | 
| 32521 | 441  | 
end  | 
| 
37994
 
b04307085a09
make TPTP generator accept full first-order formulas
 
blanchet 
parents: 
37631 
diff
changeset
 | 
442  | 
  handle ERROR msg => ("error: " ^ msg, SH_ERROR)
 | 
| 32521 | 443  | 
|
| 
32454
 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 
boehmes 
parents: 
32452 
diff
changeset
 | 
444  | 
fun thms_of_name ctxt name =  | 
| 
 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 
boehmes 
parents: 
32452 
diff
changeset
 | 
445  | 
let  | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36959 
diff
changeset
 | 
446  | 
val lex = Keyword.get_lexicons  | 
| 42361 | 447  | 
val get = maps (Proof_Context.get_fact ctxt o fst)  | 
| 
32454
 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 
boehmes 
parents: 
32452 
diff
changeset
 | 
448  | 
in  | 
| 
 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 
boehmes 
parents: 
32452 
diff
changeset
 | 
449  | 
Source.of_string name  | 
| 40526 | 450  | 
|> Symbol.source  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
451  | 
    |> Token.source {do_recover=SOME false} lex Position.start
 | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
452  | 
|> Token.source_proper  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36954 
diff
changeset
 | 
453  | 
|> Source.source Token.stopper (Parse_Spec.xthms1 >> get) NONE  | 
| 
32454
 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 
boehmes 
parents: 
32452 
diff
changeset
 | 
454  | 
|> Source.exhaust  | 
| 
 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 
boehmes 
parents: 
32452 
diff
changeset
 | 
455  | 
end  | 
| 
32452
 
d84edd022efe
apply metis with found theorems in case sledgehammer was successful
 
boehmes 
parents: 
32434 
diff
changeset
 | 
456  | 
|
| 
32498
 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 
boehmes 
parents: 
32496 
diff
changeset
 | 
457  | 
in  | 
| 
 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 
boehmes 
parents: 
32496 
diff
changeset
 | 
458  | 
|
| 44090 | 459  | 
fun run_sledgehammer trivial args reconstructor named_thms id  | 
460  | 
      ({pre=st, log, pos, ...}: Mirabelle.run_args) =
 | 
|
| 
32385
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
461  | 
let  | 
| 39340 | 462  | 
val triv_str = if trivial then "[T] " else ""  | 
| 32536 | 463  | 
val _ = change_data id inc_sh_calls  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
464  | 
val _ = if trivial then () else change_data id inc_sh_nontriv_calls  | 
| 
42444
 
8e5438dc70bb
cleanup: get rid of "may_slice" arguments without changing semantics
 
blanchet 
parents: 
42443 
diff
changeset
 | 
465  | 
val (prover_name, prover) = get_prover (Proof.context_of st) args  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43569 
diff
changeset
 | 
466  | 
val type_enc = AList.lookup (op =) args type_encK |> the_default "smart"  | 
| 44089 | 467  | 
val sound = AList.lookup (op =) args soundK |> the_default "false"  | 
| 41752 | 468  | 
val max_relevant = AList.lookup (op =) args max_relevantK |> the_default "smart"  | 
| 
42725
 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 
blanchet 
parents: 
42642 
diff
changeset
 | 
469  | 
val slicing = AList.lookup (op =) args slicingK |> the_default "true"  | 
| 
43827
 
62d64709af3b
added option to control which lambda translation to use (for experiments)
 
blanchet 
parents: 
43710 
diff
changeset
 | 
470  | 
val lambda_translation = AList.lookup (op =) args lambda_translationK  | 
| 
42725
 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 
blanchet 
parents: 
42642 
diff
changeset
 | 
471  | 
val e_weight_method = AList.lookup (op =) args e_weight_methodK  | 
| 44099 | 472  | 
val force_sos = AList.lookup (op =) args force_sosK  | 
| 
42725
 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 
blanchet 
parents: 
42642 
diff
changeset
 | 
473  | 
|> Option.map (curry (op <>) "false")  | 
| 32525 | 474  | 
val dir = AList.lookup (op =) args keepK  | 
| 32541 | 475  | 
val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30)  | 
| 41268 | 476  | 
(* always use a hard timeout, but give some slack so that the automatic  | 
477  | 
minimizer has a chance to do its magic *)  | 
|
478  | 
val hard_timeout = SOME (2 * timeout)  | 
|
| 
41155
 
85da8cbb4966
added support for "type_sys" option to Mirabelle
 
blanchet 
parents: 
41154 
diff
changeset
 | 
479  | 
val (msg, result) =  | 
| 44089 | 480  | 
run_sh prover_name prover type_enc sound max_relevant slicing  | 
| 44099 | 481  | 
lambda_translation e_weight_method force_sos hard_timeout timeout dir  | 
482  | 
pos st  | 
|
| 32525 | 483  | 
in  | 
| 32536 | 484  | 
case result of  | 
| 40062 | 485  | 
SH_OK (time_isa, time_prover, names) =>  | 
| 38700 | 486  | 
let  | 
| 
44487
 
0989d8deab69
include chained facts for minimizer, otherwise it won't work
 
blanchet 
parents: 
44462 
diff
changeset
 | 
487  | 
fun get_thms (name, loc) =  | 
| 
 
0989d8deab69
include chained facts for minimizer, otherwise it won't work
 
blanchet 
parents: 
44462 
diff
changeset
 | 
488  | 
SOME ((name, loc), thms_of_name (Proof.context_of st) name)  | 
| 32525 | 489  | 
in  | 
| 32818 | 490  | 
change_data id inc_sh_success;  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
491  | 
if trivial then () else change_data id inc_sh_nontriv_success;  | 
| 32818 | 492  | 
change_data id (inc_sh_lemmas (length names));  | 
493  | 
change_data id (inc_sh_max_lems (length names));  | 
|
494  | 
change_data id (inc_sh_time_isa time_isa);  | 
|
| 40062 | 495  | 
change_data id (inc_sh_time_prover time_prover);  | 
| 
41357
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
496  | 
reconstructor := reconstructor_from_msg args msg;  | 
| 38826 | 497  | 
named_thms := SOME (map_filter get_thms names);  | 
| 39340 | 498  | 
          log (sh_tag id ^ triv_str ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
 | 
| 40062 | 499  | 
string_of_int time_prover ^ ") [" ^ prover_name ^ "]:\n" ^ msg)  | 
| 32525 | 500  | 
end  | 
| 40062 | 501  | 
| SH_FAIL (time_isa, time_prover) =>  | 
| 32536 | 502  | 
let  | 
503  | 
val _ = change_data id (inc_sh_time_isa time_isa)  | 
|
| 40062 | 504  | 
val _ = change_data id (inc_sh_time_prover_fail time_prover)  | 
| 39340 | 505  | 
in log (sh_tag id ^ triv_str ^ "failed: " ^ msg) end  | 
| 32536 | 506  | 
| SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)  | 
| 32525 | 507  | 
end  | 
508  | 
||
509  | 
end  | 
|
510  | 
||
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
511  | 
fun run_minimize args reconstructor named_thms id  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
512  | 
        ({pre=st, log, ...}: Mirabelle.run_args) =
 | 
| 32525 | 513  | 
let  | 
| 40069 | 514  | 
val ctxt = Proof.context_of st  | 
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
515  | 
val n0 = length (these (!named_thms))  | 
| 
42444
 
8e5438dc70bb
cleanup: get rid of "may_slice" arguments without changing semantics
 
blanchet 
parents: 
42443 
diff
changeset
 | 
516  | 
val (prover_name, _) = get_prover ctxt args  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43569 
diff
changeset
 | 
517  | 
val type_enc = AList.lookup (op =) args type_encK |> the_default "smart"  | 
| 44089 | 518  | 
val sound = AList.lookup (op =) args soundK |> the_default "false"  | 
| 32525 | 519  | 
val timeout =  | 
520  | 
AList.lookup (op =) args minimize_timeoutK  | 
|
| 
40627
 
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
 
wenzelm 
parents: 
40554 
diff
changeset
 | 
521  | 
|> Option.map (fst o read_int o raw_explode) (* FIXME Symbol.explode (?) *)  | 
| 32525 | 522  | 
|> the_default 5  | 
| 
43064
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43052 
diff
changeset
 | 
523  | 
val params = Sledgehammer_Isar.default_params ctxt  | 
| 
41155
 
85da8cbb4966
added support for "type_sys" option to Mirabelle
 
blanchet 
parents: 
41154 
diff
changeset
 | 
524  | 
      [("provers", prover_name),
 | 
| 
 
85da8cbb4966
added support for "type_sys" option to Mirabelle
 
blanchet 
parents: 
41154 
diff
changeset
 | 
525  | 
       ("verbose", "true"),
 | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43569 
diff
changeset
 | 
526  | 
       ("type_enc", type_enc),
 | 
| 44089 | 527  | 
       ("sound", sound),
 | 
| 
44448
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
528  | 
       ("timeout", string_of_int timeout),
 | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
529  | 
       ("preplay_timeout", preplay_timeout)]
 | 
| 37587 | 530  | 
val minimize =  | 
| 
41742
 
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
 
blanchet 
parents: 
41741 
diff
changeset
 | 
531  | 
Sledgehammer_Minimize.minimize_facts prover_name params  | 
| 
43064
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43052 
diff
changeset
 | 
532  | 
true 1 (Sledgehammer_Util.subgoal_count st)  | 
| 32525 | 533  | 
val _ = log separator  | 
| 
43261
 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 
blanchet 
parents: 
43228 
diff
changeset
 | 
534  | 
val (used_facts, (preplay, message, message_tail)) =  | 
| 
 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 
blanchet 
parents: 
43228 
diff
changeset
 | 
535  | 
minimize st (these (!named_thms))  | 
| 
 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 
blanchet 
parents: 
43228 
diff
changeset
 | 
536  | 
val msg = message (preplay ()) ^ message_tail  | 
| 32525 | 537  | 
in  | 
| 
43052
 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 
blanchet 
parents: 
43051 
diff
changeset
 | 
538  | 
case used_facts of  | 
| 
 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 
blanchet 
parents: 
43051 
diff
changeset
 | 
539  | 
SOME named_thms' =>  | 
| 32609 | 540  | 
(change_data id inc_min_succs;  | 
541  | 
change_data id (inc_min_ab_ratios ((100 * length named_thms') div n0));  | 
|
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
542  | 
if length named_thms' = n0  | 
| 
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
543  | 
then log (minimize_tag id ^ "already minimal")  | 
| 
41357
 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 
blanchet 
parents: 
41338 
diff
changeset
 | 
544  | 
else (reconstructor := reconstructor_from_msg args msg;  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
545  | 
named_thms := SOME named_thms';  | 
| 
32571
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
546  | 
log (minimize_tag id ^ "succeeded:\n" ^ msg))  | 
| 
 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 
nipkow 
parents: 
32567 
diff
changeset
 | 
547  | 
)  | 
| 
43052
 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 
blanchet 
parents: 
43051 
diff
changeset
 | 
548  | 
| NONE => log (minimize_tag id ^ "failed: " ^ msg)  | 
| 32525 | 549  | 
end  | 
550  | 
||
551  | 
||
| 44461 | 552  | 
fun e_override_params timeout =  | 
| 44430 | 553  | 
  [("provers", "e"),
 | 
| 
44449
 
b622a98b79fb
don't select facts when using sledgehammer_tac for reconstruction
 
blanchet 
parents: 
44448 
diff
changeset
 | 
554  | 
   ("max_relevant", "0"),
 | 
| 44430 | 555  | 
   ("type_enc", "poly_guards?"),
 | 
556  | 
   ("sound", "true"),
 | 
|
| 44461 | 557  | 
   ("slicing", "false"),
 | 
558  | 
   ("timeout", timeout |> Time.toSeconds |> string_of_int)]
 | 
|
| 44430 | 559  | 
|
| 44461 | 560  | 
fun vampire_override_params timeout =  | 
| 44430 | 561  | 
  [("provers", "vampire"),
 | 
| 
44449
 
b622a98b79fb
don't select facts when using sledgehammer_tac for reconstruction
 
blanchet 
parents: 
44448 
diff
changeset
 | 
562  | 
   ("max_relevant", "0"),
 | 
| 44430 | 563  | 
   ("type_enc", "poly_tags"),
 | 
564  | 
   ("sound", "true"),
 | 
|
| 44461 | 565  | 
   ("slicing", "false"),
 | 
566  | 
   ("timeout", timeout |> Time.toSeconds |> string_of_int)]
 | 
|
| 44430 | 567  | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
568  | 
fun run_reconstructor trivial full m name reconstructor named_thms id  | 
| 
32567
 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 
wenzelm 
parents: 
32564 
diff
changeset
 | 
569  | 
    ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
 | 
| 32525 | 570  | 
let  | 
| 
44462
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
571  | 
fun do_reconstructor named_thms ctxt =  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
572  | 
let  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
573  | 
val ref_of_str =  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
574  | 
suffix ";" #> Outer_Syntax.scan Position.none #> Parse_Spec.xthm  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
575  | 
#> fst  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
576  | 
val thms = named_thms |> maps snd  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
577  | 
val facts = named_thms |> map (ref_of_str o fst o fst)  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
578  | 
        val relevance_override = {add = facts, del = [], only = true}
 | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
579  | 
in  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
580  | 
if !reconstructor = "sledgehammer_tac" then  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
581  | 
Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
582  | 
(e_override_params timeout) relevance_override  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
583  | 
ORELSE'  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
584  | 
Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
585  | 
(vampire_override_params timeout) relevance_override  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
586  | 
else if !reconstructor = "smt" then  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
587  | 
SMT_Solver.smt_tac ctxt thms  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
588  | 
else if full orelse !reconstructor = "metis (full_types)" then  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
589  | 
Metis_Tactics.metis_tac [Metis_Tactics.full_type_enc] ctxt thms  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
590  | 
else if !reconstructor = "metis (no_types)" then  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
591  | 
Metis_Tactics.metis_tac [Metis_Tactics.no_type_enc] ctxt thms  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
592  | 
else if !reconstructor = "metis" then  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
593  | 
Metis_Tactics.metis_tac [] ctxt thms  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
594  | 
else  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
595  | 
K all_tac  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
596  | 
end  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
597  | 
fun apply_reconstructor named_thms =  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
598  | 
Mirabelle.can_apply timeout (do_reconstructor named_thms) st  | 
| 32521 | 599  | 
|
600  | 
    fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
 | 
|
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
601  | 
| with_time (true, t) = (change_data id (inc_reconstructor_success m);  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
602  | 
if trivial then ()  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
603  | 
else change_data id (inc_reconstructor_nontriv_success m);  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
604  | 
change_data id (inc_reconstructor_lemmas m (length named_thms));  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
605  | 
change_data id (inc_reconstructor_time m t);  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
606  | 
change_data id (inc_reconstructor_posns m (pos, trivial));  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
607  | 
if name = "proof" then change_data id (inc_reconstructor_proofs m)  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
608  | 
else ();  | 
| 32521 | 609  | 
          "succeeded (" ^ string_of_int t ^ ")")
 | 
| 
44462
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
610  | 
fun timed_reconstructor named_thms =  | 
| 
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
611  | 
(with_time (Mirabelle.cpu_time apply_reconstructor named_thms), true)  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
612  | 
handle TimeLimit.TimeOut => (change_data id (inc_reconstructor_timeout m);  | 
| 34052 | 613  | 
               ("timeout", false))
 | 
614  | 
           | ERROR msg => ("error: " ^ msg, false)
 | 
|
| 32521 | 615  | 
|
| 32525 | 616  | 
val _ = log separator  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
617  | 
val _ = change_data id (inc_reconstructor_calls m)  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
618  | 
val _ = if trivial then ()  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
619  | 
else change_data id (inc_reconstructor_nontriv_calls m)  | 
| 32521 | 620  | 
in  | 
| 
44462
 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 
blanchet 
parents: 
44461 
diff
changeset
 | 
621  | 
named_thms  | 
| 
40667
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
622  | 
|> timed_reconstructor  | 
| 
 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 
blanchet 
parents: 
40627 
diff
changeset
 | 
623  | 
|>> log o prefix (reconstructor_tag reconstructor id)  | 
| 34052 | 624  | 
|> snd  | 
| 32521 | 625  | 
end  | 
| 
32385
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
626  | 
|
| 
41276
 
285aea0c153c
two layers of timeouts seem to be less reliable than a single layer
 
blanchet 
parents: 
41275 
diff
changeset
 | 
627  | 
val try_timeout = seconds 5.0  | 
| 
39337
 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 
blanchet 
parents: 
39321 
diff
changeset
 | 
628  | 
|
| 
44431
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
629  | 
(* crude hack *)  | 
| 
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
630  | 
val num_sledgehammer_calls = Unsynchronized.ref 0  | 
| 
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
631  | 
|
| 
34035
 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 
boehmes 
parents: 
33316 
diff
changeset
 | 
632  | 
fun sledgehammer_action args id (st as {pre, name, ...}: Mirabelle.run_args) =
 | 
| 
35592
 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 
wenzelm 
parents: 
34052 
diff
changeset
 | 
633  | 
let val goal = Thm.major_prem_of (#goal (Proof.goal pre)) in  | 
| 
 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 
wenzelm 
parents: 
34052 
diff
changeset
 | 
634  | 
if can Logic.dest_conjunction goal orelse can Logic.dest_equals goal  | 
| 
 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 
wenzelm 
parents: 
34052 
diff
changeset
 | 
635  | 
then () else  | 
| 
 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 
wenzelm 
parents: 
34052 
diff
changeset
 | 
636  | 
let  | 
| 
44431
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
637  | 
val max_calls =  | 
| 
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
638  | 
AList.lookup (op =) args max_callsK |> the_default "10000000"  | 
| 
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
639  | 
|> Int.fromString |> the  | 
| 
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
640  | 
val _ = num_sledgehammer_calls := !num_sledgehammer_calls + 1;  | 
| 
44448
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
641  | 
in  | 
| 
44431
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
642  | 
if !num_sledgehammer_calls > max_calls then ()  | 
| 
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
643  | 
else  | 
| 
44448
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
644  | 
let  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
645  | 
val reconstructor = Unsynchronized.ref ""  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
646  | 
val named_thms =  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
647  | 
Unsynchronized.ref (NONE : ((string * locality) * thm list) list option)  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
648  | 
val minimize = AList.defined (op =) args minimizeK  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
649  | 
val metis_ft = AList.defined (op =) args metis_ftK  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
650  | 
val trivial =  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
651  | 
Try_Methods.try_methods (SOME try_timeout) ([], [], [], []) pre  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
652  | 
handle TimeLimit.TimeOut => false  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
653  | 
fun apply_reconstructor m1 m2 =  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
654  | 
if metis_ft  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
655  | 
then  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
656  | 
if not (Mirabelle.catch_result (reconstructor_tag reconstructor) false  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
657  | 
(run_reconstructor trivial false m1 name reconstructor  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
658  | 
(these (!named_thms))) id st)  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
659  | 
then  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
660  | 
(Mirabelle.catch_result (reconstructor_tag reconstructor) false  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
661  | 
(run_reconstructor trivial true m2 name reconstructor  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
662  | 
(these (!named_thms))) id st; ())  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
663  | 
else ()  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
664  | 
else  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
665  | 
(Mirabelle.catch_result (reconstructor_tag reconstructor) false  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
666  | 
(run_reconstructor trivial false m1 name reconstructor  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
667  | 
(these (!named_thms))) id st; ())  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
668  | 
in  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
669  | 
change_data id (set_mini minimize);  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
670  | 
Mirabelle.catch sh_tag (run_sledgehammer trivial args reconstructor  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
671  | 
named_thms) id st;  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
672  | 
if is_some (!named_thms)  | 
| 
44431
 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 
blanchet 
parents: 
44430 
diff
changeset
 | 
673  | 
then  | 
| 
44448
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
674  | 
(apply_reconstructor Unminimized UnminimizedFT;  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
675  | 
if minimize andalso not (null (these (!named_thms)))  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
676  | 
then  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
677  | 
(Mirabelle.catch minimize_tag  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
678  | 
(run_minimize args reconstructor named_thms) id st;  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
679  | 
apply_reconstructor Minimized MinimizedFT)  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
680  | 
else ())  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
681  | 
else ()  | 
| 
 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 
blanchet 
parents: 
44447 
diff
changeset
 | 
682  | 
end  | 
| 
35592
 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 
wenzelm 
parents: 
34052 
diff
changeset
 | 
683  | 
end  | 
| 32818 | 684  | 
end  | 
| 
32385
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
685  | 
|
| 32511 | 686  | 
fun invoke args =  | 
| 
43569
 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 
blanchet 
parents: 
43351 
diff
changeset
 | 
687  | 
Mirabelle.register (init, sledgehammer_action args, done)  | 
| 
32385
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
688  | 
|
| 
 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 
boehmes 
parents:  
diff
changeset
 | 
689  | 
end  |