author | wenzelm |
Sat, 13 Nov 2010 19:55:45 +0100 | |
changeset 40526 | ca3c6b1bfcdb |
parent 40417 | a29b2fee592b |
child 40554 | ff446d5e9a62 |
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" |
11 |
val full_typesK = "full_types" |
|
32525 | 12 |
val minimizeK = "minimize" |
13 |
val minimize_timeoutK = "minimize_timeout" |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
14 |
val metis_ftK = "metis_ft" |
32521 | 15 |
|
16 |
fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: " |
|
32525 | 17 |
fun minimize_tag id = "#" ^ string_of_int id ^ " minimize (sledgehammer): " |
40416
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
18 |
fun metis_tag smt id = |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
19 |
"#" ^ string_of_int id ^ " " ^ (if smt then "smt" else "metis") ^ |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
20 |
" (sledgehammer): " |
32521 | 21 |
|
32525 | 22 |
val separator = "-----" |
23 |
||
32521 | 24 |
|
32549 | 25 |
datatype sh_data = ShData of { |
26 |
calls: int, |
|
27 |
success: int, |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
28 |
nontriv_calls: int, |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
29 |
nontriv_success: int, |
32585 | 30 |
lemmas: int, |
32818 | 31 |
max_lems: int, |
32549 | 32 |
time_isa: int, |
40062 | 33 |
time_prover: int, |
34 |
time_prover_fail: int} |
|
32549 | 35 |
|
36 |
datatype me_data = MeData of { |
|
37 |
calls: int, |
|
38 |
success: int, |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
39 |
nontriv_calls: int, |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
40 |
nontriv_success: int, |
32676 | 41 |
proofs: int, |
32549 | 42 |
time: int, |
32550 | 43 |
timeout: int, |
32990 | 44 |
lemmas: int * int * int, |
39341
d2b981a0429a
indicate triviality in the list of proved things
blanchet
parents:
39340
diff
changeset
|
45 |
posns: (Position.T * bool) list |
32550 | 46 |
} |
32549 | 47 |
|
32571
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents:
32567
diff
changeset
|
48 |
datatype min_data = MinData of { |
32609 | 49 |
succs: int, |
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35830
diff
changeset
|
50 |
ab_ratios: int |
32571
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents:
32567
diff
changeset
|
51 |
} |
32521 | 52 |
|
32818 | 53 |
fun make_sh_data |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
54 |
(calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa, |
40062 | 55 |
time_prover,time_prover_fail) = |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
56 |
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
|
57 |
nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems, |
40062 | 58 |
time_isa=time_isa, time_prover=time_prover, |
59 |
time_prover_fail=time_prover_fail} |
|
32521 | 60 |
|
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35830
diff
changeset
|
61 |
fun make_min_data (succs, ab_ratios) = |
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35830
diff
changeset
|
62 |
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
|
63 |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
64 |
fun make_me_data (calls,success,nontriv_calls,nontriv_success,proofs,time, |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
65 |
timeout,lemmas,posns) = |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
66 |
MeData{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
|
67 |
nontriv_success=nontriv_success, proofs=proofs, time=time, |
32990 | 68 |
timeout=timeout, lemmas=lemmas, posns=posns} |
32549 | 69 |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
70 |
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
|
71 |
val empty_min_data = make_min_data (0, 0) |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
72 |
val empty_me_data = make_me_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
|
73 |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
74 |
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
|
75 |
lemmas, max_lems, time_isa, |
40062 | 76 |
time_prover, time_prover_fail}) = (calls, success, nontriv_calls, |
77 |
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
|
78 |
|
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35830
diff
changeset
|
79 |
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
|
80 |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
81 |
fun tuple_of_me_data (MeData {calls, success, nontriv_calls, nontriv_success, |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
82 |
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
|
83 |
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
|
84 |
|
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
85 |
|
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
86 |
datatype metis = Unminimized | Minimized | UnminimizedFT | MinimizedFT |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
87 |
|
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
88 |
datatype data = Data of { |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
89 |
sh: sh_data, |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
90 |
min: min_data, |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
91 |
me_u: me_data, (* metis with unminimized set of lemmas *) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
92 |
me_m: me_data, (* metis with minimized set of lemmas *) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
93 |
me_uft: me_data, (* metis with unminimized set of lemmas and fully-typed *) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
94 |
me_mft: me_data, (* metis with minimized set of lemmas and fully-typed *) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
95 |
mini: bool (* with minimization *) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
96 |
} |
32521 | 97 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
98 |
fun make_data (sh, min, me_u, me_m, me_uft, me_mft, mini) = |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
99 |
Data {sh=sh, min=min, me_u=me_u, me_m=me_m, me_uft=me_uft, me_mft=me_mft, |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
100 |
mini=mini} |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
101 |
|
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
102 |
val empty_data = make_data (empty_sh_data, empty_min_data, |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
103 |
empty_me_data, empty_me_data, empty_me_data, empty_me_data, false) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
104 |
|
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
105 |
fun map_sh_data f (Data {sh, min, me_u, me_m, me_uft, me_mft, mini}) = |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
106 |
let val sh' = make_sh_data (f (tuple_of_sh_data sh)) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
107 |
in make_data (sh', min, me_u, me_m, me_uft, me_mft, mini) end |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
108 |
|
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
109 |
fun map_min_data f (Data {sh, min, me_u, me_m, me_uft, me_mft, mini}) = |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
110 |
let val min' = make_min_data (f (tuple_of_min_data min)) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
111 |
in make_data (sh, min', me_u, me_m, me_uft, me_mft, mini) end |
32521 | 112 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
113 |
fun map_me_data f m (Data {sh, min, me_u, me_m, me_uft, me_mft, mini}) = |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
114 |
let |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
115 |
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
|
116 |
| 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
|
117 |
| 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
|
118 |
| map_me g MinimizedFT (u, m, uft, mft) = (u, m, uft, g mft) |
32521 | 119 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
120 |
val f' = make_me_data o f o tuple_of_me_data |
32571
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents:
32567
diff
changeset
|
121 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
122 |
val (me_u', me_m', me_uft', me_mft') = |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
123 |
map_me f' m (me_u, me_m, me_uft, me_mft) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
124 |
in make_data (sh, min, me_u', me_m', me_uft', me_mft', mini) end |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
125 |
|
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
126 |
fun set_mini mini (Data {sh, min, me_u, me_m, me_uft, me_mft, ...}) = |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
127 |
make_data (sh, min, me_u, me_m, me_uft, me_mft, mini) |
32990 | 128 |
|
129 |
fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n)); |
|
32521 | 130 |
|
32818 | 131 |
val inc_sh_calls = map_sh_data |
40062 | 132 |
(fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail) |
133 |
=> (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail)) |
|
32549 | 134 |
|
32818 | 135 |
val inc_sh_success = map_sh_data |
40062 | 136 |
(fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail) |
137 |
=> (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
|
138 |
|
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
139 |
val inc_sh_nontriv_calls = map_sh_data |
40062 | 140 |
(fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail) |
141 |
=> (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
|
142 |
|
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
143 |
val inc_sh_nontriv_success = map_sh_data |
40062 | 144 |
(fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail) |
145 |
=> (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_prover, time_prover_fail)) |
|
32585 | 146 |
|
32818 | 147 |
fun inc_sh_lemmas n = map_sh_data |
40062 | 148 |
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) |
149 |
=> (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_prover,time_prover_fail)) |
|
32521 | 150 |
|
32818 | 151 |
fun inc_sh_max_lems n = map_sh_data |
40062 | 152 |
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) |
153 |
=> (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_prover,time_prover_fail)) |
|
32549 | 154 |
|
32818 | 155 |
fun inc_sh_time_isa t = map_sh_data |
40062 | 156 |
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) |
157 |
=> (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_prover,time_prover_fail)) |
|
32818 | 158 |
|
40062 | 159 |
fun inc_sh_time_prover t = map_sh_data |
160 |
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) |
|
161 |
=> (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover + t,time_prover_fail)) |
|
32521 | 162 |
|
40062 | 163 |
fun inc_sh_time_prover_fail t = map_sh_data |
164 |
(fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) |
|
165 |
=> (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
|
166 |
|
32818 | 167 |
val inc_min_succs = map_min_data |
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35830
diff
changeset
|
168 |
(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
|
169 |
|
32818 | 170 |
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
|
171 |
(fn (succs, ab_ratios) => (succs, ab_ratios+r)) |
32549 | 172 |
|
32551
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
173 |
val inc_metis_calls = map_me_data |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
174 |
(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
|
175 |
=> (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns)) |
32533 | 176 |
|
32551
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
177 |
val inc_metis_success = map_me_data |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
178 |
(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
|
179 |
=> (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
|
180 |
|
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
181 |
val inc_metis_nontriv_calls = map_me_data |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
182 |
(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
|
183 |
=> (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
|
184 |
|
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
185 |
val inc_metis_nontriv_success = map_me_data |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
186 |
(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
|
187 |
=> (calls, success, nontriv_calls, nontriv_success + 1, proofs, time, timeout, lemmas,posns)) |
32676 | 188 |
|
189 |
val inc_metis_proofs = map_me_data |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
190 |
(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
|
191 |
=> (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
|
192 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
193 |
fun inc_metis_time m t = map_me_data |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
194 |
(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
|
195 |
=> (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns)) m |
32536 | 196 |
|
32551
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
197 |
val inc_metis_timeout = map_me_data |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
198 |
(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
|
199 |
=> (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout + 1, lemmas,posns)) |
32549 | 200 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
201 |
fun inc_metis_lemmas m n = map_me_data |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
202 |
(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
|
203 |
=> (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
|
204 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
205 |
fun inc_metis_posns m pos = map_me_data |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
206 |
(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
|
207 |
=> (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns)) m |
32521 | 208 |
|
209 |
local |
|
210 |
||
211 |
val str = string_of_int |
|
212 |
val str3 = Real.fmt (StringCvt.FIX (SOME 3)) |
|
213 |
fun percentage a b = string_of_int (a * 100 div b) |
|
214 |
fun time t = Real.fromInt t / 1000.0 |
|
215 |
fun avg_time t n = |
|
216 |
if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0 |
|
217 |
||
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
218 |
fun log_sh_data log |
40062 | 219 |
(calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail) = |
32818 | 220 |
(log ("Total number of sledgehammer calls: " ^ str calls); |
221 |
log ("Number of successful sledgehammer calls: " ^ str success); |
|
222 |
log ("Number of sledgehammer lemmas: " ^ str lemmas); |
|
223 |
log ("Max number of sledgehammer lemmas: " ^ str max_lems); |
|
224 |
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
|
225 |
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
|
226 |
log ("Number of successful nontrivial sledgehammer calls: " ^ str nontriv_success); |
32818 | 227 |
log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa)); |
40062 | 228 |
log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_prover)); |
229 |
log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_prover_fail)); |
|
32536 | 230 |
log ("Average time for sledgehammer calls (Isabelle): " ^ |
32818 | 231 |
str3 (avg_time time_isa calls)); |
32533 | 232 |
log ("Average time for successful sledgehammer calls (ATP): " ^ |
40062 | 233 |
str3 (avg_time time_prover success)); |
32536 | 234 |
log ("Average time for failed sledgehammer calls (ATP): " ^ |
40062 | 235 |
str3 (avg_time time_prover_fail (calls - success))) |
32533 | 236 |
) |
32521 | 237 |
|
32551
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
238 |
|
39341
d2b981a0429a
indicate triviality in the list of proved things
blanchet
parents:
39340
diff
changeset
|
239 |
fun str_of_pos (pos, triv) = |
32551
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
240 |
let val str0 = string_of_int o the_default 0 |
39341
d2b981a0429a
indicate triviality in the list of proved things
blanchet
parents:
39340
diff
changeset
|
241 |
in |
d2b981a0429a
indicate triviality in the list of proved things
blanchet
parents:
39340
diff
changeset
|
242 |
str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) ^ |
d2b981a0429a
indicate triviality in the list of proved things
blanchet
parents:
39340
diff
changeset
|
243 |
(if triv then "[T]" else "") |
d2b981a0429a
indicate triviality in the list of proved things
blanchet
parents:
39340
diff
changeset
|
244 |
end |
32551
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
245 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
246 |
fun log_me_data log tag sh_calls (metis_calls, metis_success, |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
247 |
metis_nontriv_calls, metis_nontriv_success, |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
248 |
metis_proofs, metis_time, metis_timeout, (lemmas, lems_sos, lems_max), |
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
249 |
metis_posns) = |
32549 | 250 |
(log ("Total number of " ^ tag ^ "metis calls: " ^ str metis_calls); |
32740 | 251 |
log ("Number of successful " ^ tag ^ "metis calls: " ^ str metis_success ^ |
252 |
" (proof: " ^ str metis_proofs ^ ")"); |
|
32549 | 253 |
log ("Number of " ^ tag ^ "metis timeouts: " ^ str metis_timeout); |
32533 | 254 |
log ("Success rate: " ^ percentage metis_success sh_calls ^ "%"); |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
255 |
log ("Total number of nontrivial " ^ tag ^ "metis calls: " ^ str metis_nontriv_calls); |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
256 |
log ("Number of successful nontrivial " ^ tag ^ "metis calls: " ^ str metis_nontriv_success ^ |
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
257 |
" (proof: " ^ str metis_proofs ^ ")"); |
32990 | 258 |
log ("Number of successful " ^ tag ^ "metis lemmas: " ^ str lemmas); |
259 |
log ("SOS of successful " ^ tag ^ "metis lemmas: " ^ str lems_sos); |
|
260 |
log ("Max number of successful " ^ tag ^ "metis lemmas: " ^ str lems_max); |
|
261 |
log ("Total time for successful " ^ tag ^ "metis calls: " ^ str3 (time metis_time)); |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
262 |
log ("Average time for successful " ^ tag ^ "metis calls: " ^ |
32551
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
263 |
str3 (avg_time metis_time metis_success)); |
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
264 |
if tag="" |
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
265 |
then log ("Proved: " ^ space_implode " " (map str_of_pos metis_posns)) |
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
266 |
else () |
421323205efd
position information is now passed to all actions;
nipkow
parents:
32550
diff
changeset
|
267 |
) |
32571
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents:
32567
diff
changeset
|
268 |
|
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35830
diff
changeset
|
269 |
fun log_min_data log (succs, ab_ratios) = |
32609 | 270 |
(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
|
271 |
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
|
272 |
) |
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents:
32567
diff
changeset
|
273 |
|
32521 | 274 |
in |
275 |
||
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
276 |
fun log_data id log (Data {sh, min, me_u, me_m, me_uft, me_mft, mini}) = |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
277 |
let |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
278 |
val ShData {calls=sh_calls, ...} = sh |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
279 |
|
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
280 |
fun app_if (MeData {calls, ...}) f = if calls > 0 then f () else () |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
281 |
fun log_me tag m = |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
282 |
log_me_data log tag sh_calls (tuple_of_me_data m) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
283 |
fun log_metis (tag1, m1) (tag2, m2) = app_if m1 (fn () => |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
284 |
(log_me tag1 m1; log ""; app_if m2 (fn () => log_me tag2 m2))) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
285 |
in |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
286 |
if sh_calls > 0 |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
287 |
then |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
288 |
(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
|
289 |
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
|
290 |
log ""; |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
291 |
if not mini |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
292 |
then log_metis ("", me_u) ("fully-typed ", me_uft) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
293 |
else |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
294 |
app_if me_u (fn () => |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
295 |
(log_metis ("unminimized ", me_u) ("unminimized fully-typed ", me_uft); |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
296 |
log ""; |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
297 |
app_if me_m (fn () => |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
298 |
(log_min_data log (tuple_of_min_data min); log ""; |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
299 |
log_metis ("", me_m) ("fully-typed ", me_mft)))))) |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
300 |
else () |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
301 |
end |
32521 | 302 |
|
303 |
end |
|
304 |
||
305 |
||
306 |
(* Warning: we implicitly assume single-threaded execution here! *) |
|
32740 | 307 |
val data = Unsynchronized.ref ([] : (int * data) list) |
32521 | 308 |
|
32740 | 309 |
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
|
310 |
fun done id ({log, ...}: Mirabelle.done_args) = |
32521 | 311 |
AList.lookup (op =) (!data) id |
312 |
|> Option.map (log_data id log) |
|
313 |
|> K () |
|
314 |
||
32740 | 315 |
fun change_data id f = (Unsynchronized.change data (AList.map_entry (op =) id f); ()) |
32521 | 316 |
|
317 |
||
40069 | 318 |
fun get_prover ctxt args = |
33016
b73b74fe23c3
proper exceptions instead of unhandled partiality
boehmes
parents:
32991
diff
changeset
|
319 |
let |
40200 | 320 |
val thy = ProofContext.theory_of ctxt |
40062 | 321 |
fun default_prover_name () = |
40069 | 322 |
hd (#provers (Sledgehammer_Isar.default_params ctxt [])) |
33016
b73b74fe23c3
proper exceptions instead of unhandled partiality
boehmes
parents:
32991
diff
changeset
|
323 |
handle Empty => error "No ATP available." |
40200 | 324 |
fun get_prover name = (name, Sledgehammer.get_prover thy false name) |
33016
b73b74fe23c3
proper exceptions instead of unhandled partiality
boehmes
parents:
32991
diff
changeset
|
325 |
in |
b73b74fe23c3
proper exceptions instead of unhandled partiality
boehmes
parents:
32991
diff
changeset
|
326 |
(case AList.lookup (op =) args proverK of |
40062 | 327 |
SOME name => get_prover name |
328 |
| NONE => get_prover (default_prover_name ())) |
|
33016
b73b74fe23c3
proper exceptions instead of unhandled partiality
boehmes
parents:
32991
diff
changeset
|
329 |
end |
32525 | 330 |
|
38988 | 331 |
type locality = Sledgehammer_Filter.locality |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38700
diff
changeset
|
332 |
|
32521 | 333 |
local |
334 |
||
32536 | 335 |
datatype sh_result = |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38700
diff
changeset
|
336 |
SH_OK of int * int * (string * locality) list | |
32536 | 337 |
SH_FAIL of int * int | |
338 |
SH_ERROR |
|
339 |
||
38998 | 340 |
fun run_sh prover_name prover hard_timeout timeout dir st = |
32521 | 341 |
let |
38998 | 342 |
val {context = ctxt, facts = chained_ths, goal} = Proof.goal st |
35971 | 343 |
val thy = ProofContext.theory_of ctxt |
38998 | 344 |
val i = 1 |
39321 | 345 |
fun change_dir (SOME dir) = Config.put Sledgehammer.dest_dir dir |
346 |
| change_dir NONE = I |
|
347 |
val st' = |
|
348 |
st |> Proof.map_context |
|
349 |
(change_dir dir |
|
350 |
#> Config.put Sledgehammer.measure_run_time true) |
|
38998 | 351 |
val params as {full_types, relevance_thresholds, max_relevant, ...} = |
40069 | 352 |
Sledgehammer_Isar.default_params ctxt |
40372
f16ce22f34d0
remove " s" suffix since seconds are now implicit
blanchet
parents:
40369
diff
changeset
|
353 |
[("timeout", Int.toString timeout)] |
40062 | 354 |
val default_max_relevant = |
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
355 |
Sledgehammer.default_max_relevant_for_prover thy prover_name |
40369
53dca3bd4250
use the SMT integration's official list of built-ins
blanchet
parents:
40301
diff
changeset
|
356 |
val is_built_in_const = |
53dca3bd4250
use the SMT integration's official list of built-ins
blanchet
parents:
40301
diff
changeset
|
357 |
Sledgehammer.is_built_in_const_for_prover ctxt prover_name |
40070
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
358 |
val relevance_fudge = Sledgehammer.relevance_fudge_for_prover prover_name |
bdb890782d4a
replaced references with proper record that's threaded through
blanchet
parents:
40069
diff
changeset
|
359 |
val relevance_override = {add = [], del = [], only = false} |
38998 | 360 |
val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal i |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40200
diff
changeset
|
361 |
val facts = |
38998 | 362 |
Sledgehammer_Filter.relevant_facts ctxt full_types relevance_thresholds |
40369
53dca3bd4250
use the SMT integration's official list of built-ins
blanchet
parents:
40301
diff
changeset
|
363 |
(the_default default_max_relevant max_relevant) is_built_in_const |
40071
658a37c80b53
generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents:
40070
diff
changeset
|
364 |
relevance_fudge relevance_override chained_ths hyp_ts concl_t |
39004
f1b465f889b5
translate the axioms to FOF once and for all ATPs
blanchet
parents:
39003
diff
changeset
|
365 |
val problem = |
39321 | 366 |
{state = st', goal = goal, subgoal = i, |
40065 | 367 |
subgoal_count = Sledgehammer_Util.subgoal_count st, |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40200
diff
changeset
|
368 |
facts = facts |> map Sledgehammer.Untranslated, only = true} |
32574
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents:
32571
diff
changeset
|
369 |
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
|
370 |
(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
|
371 |
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
|
372 |
| SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs)) |
40374
443b426e05ea
make Mirabelle work correctly if the prover (e.g. the SMT solver) returns no timing information
blanchet
parents:
40372
diff
changeset
|
373 |
val ({outcome, message, used_facts, run_time_in_msecs, ...} |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
374 |
: Sledgehammer.prover_result, |
35971 | 375 |
time_isa) = time_limit (Mirabelle.cpu_time |
38103 | 376 |
(prover params (K ""))) problem |
40374
443b426e05ea
make Mirabelle work correctly if the prover (e.g. the SMT solver) returns no timing information
blanchet
parents:
40372
diff
changeset
|
377 |
val time_prover = run_time_in_msecs |> the_default ~1 |
32521 | 378 |
in |
36405 | 379 |
case outcome of |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40200
diff
changeset
|
380 |
NONE => (message, SH_OK (time_isa, time_prover, used_facts)) |
40062 | 381 |
| SOME _ => (message, SH_FAIL (time_isa, time_prover)) |
32521 | 382 |
end |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37631
diff
changeset
|
383 |
handle ERROR msg => ("error: " ^ msg, SH_ERROR) |
32574
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents:
32571
diff
changeset
|
384 |
| TimeLimit.TimeOut => ("timeout", SH_ERROR) |
32521 | 385 |
|
32454
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
386 |
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
|
387 |
let |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36959
diff
changeset
|
388 |
val lex = Keyword.get_lexicons |
32454
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
389 |
val get = maps (ProofContext.get_fact ctxt o fst) |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
390 |
in |
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents:
32452
diff
changeset
|
391 |
Source.of_string name |
40526 | 392 |
|> 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
|
393 |
|> 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
|
394 |
|> 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
|
395 |
|> 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
|
396 |
|> 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
|
397 |
end |
32452
d84edd022efe
apply metis with found theorems in case sledgehammer was successful
boehmes
parents:
32434
diff
changeset
|
398 |
|
32498
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
399 |
in |
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
boehmes
parents:
32496
diff
changeset
|
400 |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
401 |
fun run_sledgehammer trivial args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) = |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
402 |
let |
39340 | 403 |
val triv_str = if trivial then "[T] " else "" |
32536 | 404 |
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
|
405 |
val _ = if trivial then () else change_data id inc_sh_nontriv_calls |
40069 | 406 |
val (prover_name, prover) = get_prover (Proof.context_of st) args |
32525 | 407 |
val dir = AList.lookup (op =) args keepK |
32541 | 408 |
val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30) |
40417 | 409 |
val hard_timeout = SOME timeout (* always use a hard timeout *) |
38998 | 410 |
val (msg, result) = run_sh prover_name prover hard_timeout timeout dir st |
32525 | 411 |
in |
32536 | 412 |
case result of |
40062 | 413 |
SH_OK (time_isa, time_prover, names) => |
38700 | 414 |
let |
39377 | 415 |
fun get_thms (_, Sledgehammer_Filter.Chained) = NONE |
38826 | 416 |
| get_thms (name, loc) = |
417 |
SOME ((name, loc), thms_of_name (Proof.context_of st) name) |
|
32525 | 418 |
in |
32818 | 419 |
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
|
420 |
if trivial then () else change_data id inc_sh_nontriv_success; |
32818 | 421 |
change_data id (inc_sh_lemmas (length names)); |
422 |
change_data id (inc_sh_max_lems (length names)); |
|
423 |
change_data id (inc_sh_time_isa time_isa); |
|
40062 | 424 |
change_data id (inc_sh_time_prover time_prover); |
38826 | 425 |
named_thms := SOME (map_filter get_thms names); |
39340 | 426 |
log (sh_tag id ^ triv_str ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^ |
40062 | 427 |
string_of_int time_prover ^ ") [" ^ prover_name ^ "]:\n" ^ msg) |
32525 | 428 |
end |
40062 | 429 |
| SH_FAIL (time_isa, time_prover) => |
32536 | 430 |
let |
431 |
val _ = change_data id (inc_sh_time_isa time_isa) |
|
40062 | 432 |
val _ = change_data id (inc_sh_time_prover_fail time_prover) |
39340 | 433 |
in log (sh_tag id ^ triv_str ^ "failed: " ^ msg) end |
32536 | 434 |
| SH_ERROR => log (sh_tag id ^ "failed: " ^ msg) |
32525 | 435 |
end |
436 |
||
437 |
end |
|
438 |
||
32567
de411627a985
explicitly export type abbreviations (as usual in SML97);
wenzelm
parents:
32564
diff
changeset
|
439 |
fun run_minimize args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) = |
32525 | 440 |
let |
40069 | 441 |
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
|
442 |
val n0 = length (these (!named_thms)) |
40069 | 443 |
val (prover_name, _) = get_prover ctxt args |
32525 | 444 |
val timeout = |
445 |
AList.lookup (op =) args minimize_timeoutK |
|
446 |
|> Option.map (fst o read_int o explode) |
|
447 |
|> the_default 5 |
|
40069 | 448 |
val params = Sledgehammer_Isar.default_params ctxt |
40372
f16ce22f34d0
remove " s" suffix since seconds are now implicit
blanchet
parents:
40369
diff
changeset
|
449 |
[("provers", prover_name), ("timeout", Int.toString timeout)] |
37587 | 450 |
val minimize = |
40065 | 451 |
Sledgehammer_Minimize.minimize_facts params 1 |
452 |
(Sledgehammer_Util.subgoal_count st) |
|
32525 | 453 |
val _ = log separator |
454 |
in |
|
35971 | 455 |
case minimize st (these (!named_thms)) of |
35871
c93bda4fdf15
remove the iteration counter from Sledgehammer's minimizer
blanchet
parents:
35867
diff
changeset
|
456 |
(SOME named_thms', msg) => |
32609 | 457 |
(change_data id inc_min_succs; |
458 |
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
|
459 |
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
|
460 |
then log (minimize_tag id ^ "already minimal") |
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents:
32567
diff
changeset
|
461 |
else (named_thms := SOME named_thms'; |
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents:
32567
diff
changeset
|
462 |
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
|
463 |
) |
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents:
32567
diff
changeset
|
464 |
| (NONE, msg) => log (minimize_tag id ^ "failed: " ^ msg) |
32525 | 465 |
end |
466 |
||
467 |
||
40416
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
468 |
fun run_metis smt trivial full m name named_thms id |
32567
de411627a985
explicitly export type abbreviations (as usual in SML97);
wenzelm
parents:
32564
diff
changeset
|
469 |
({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) = |
32525 | 470 |
let |
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
471 |
fun metis thms ctxt = |
40416
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
472 |
(if smt then SMT_Solver.smt_tac |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
473 |
else if full then Metis_Tactics.metisFT_tac |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
474 |
else Metis_Tactics.metis_tac) ctxt thms |
32521 | 475 |
fun apply_metis thms = Mirabelle.can_apply timeout (metis thms) st |
476 |
||
477 |
fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")" |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
478 |
| with_time (true, t) = (change_data id (inc_metis_success m); |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
479 |
if trivial then () else change_data id (inc_metis_nontriv_success m); |
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
480 |
change_data id (inc_metis_lemmas m (length named_thms)); |
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
481 |
change_data id (inc_metis_time m t); |
39341
d2b981a0429a
indicate triviality in the list of proved things
blanchet
parents:
39340
diff
changeset
|
482 |
change_data id (inc_metis_posns m (pos, trivial)); |
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
483 |
if name = "proof" then change_data id (inc_metis_proofs m) else (); |
32521 | 484 |
"succeeded (" ^ string_of_int t ^ ")") |
34052 | 485 |
fun timed_metis thms = |
486 |
(with_time (Mirabelle.cpu_time apply_metis thms), true) |
|
487 |
handle TimeLimit.TimeOut => (change_data id (inc_metis_timeout m); |
|
488 |
("timeout", false)) |
|
489 |
| ERROR msg => ("error: " ^ msg, false) |
|
32521 | 490 |
|
32525 | 491 |
val _ = log separator |
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
492 |
val _ = change_data id (inc_metis_calls m) |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
493 |
val _ = if trivial then () else change_data id (inc_metis_nontriv_calls m) |
32521 | 494 |
in |
32525 | 495 |
maps snd named_thms |
32521 | 496 |
|> timed_metis |
40416
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
497 |
|>> log o prefix (metis_tag smt id) |
34052 | 498 |
|> snd |
32521 | 499 |
end |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
500 |
|
40301 | 501 |
val try_inner_timeout = seconds 5.0 |
502 |
val try_outer_timeout = seconds 30.0 |
|
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
503 |
|
34035
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents:
33316
diff
changeset
|
504 |
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
|
505 |
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
|
506 |
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
|
507 |
then () else |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
508 |
let |
38700 | 509 |
val named_thms = |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38700
diff
changeset
|
510 |
Unsynchronized.ref (NONE : ((string * locality) * thm list) list option) |
40416
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
511 |
val ctxt = Proof.context_of pre |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
512 |
val (prover_name, _) = get_prover ctxt args |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
513 |
val smt = String.isSuffix "smt" prover_name (* hack *) |
35592
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
514 |
val minimize = AList.defined (op =) args minimizeK |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
515 |
val metis_ft = AList.defined (op =) args metis_ftK |
39363 | 516 |
val trivial = TimeLimit.timeLimit try_outer_timeout |
517 |
(Try.invoke_try (SOME try_inner_timeout)) pre |
|
518 |
handle TimeLimit.TimeOut => false |
|
35592
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
519 |
fun apply_metis m1 m2 = |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
520 |
if metis_ft |
34052 | 521 |
then |
40416
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
522 |
if not (Mirabelle.catch_result (metis_tag smt) false |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
523 |
(run_metis smt trivial false m1 name (these (!named_thms))) id st) |
35592
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
524 |
then |
40416
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
525 |
(Mirabelle.catch_result (metis_tag smt) false |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
526 |
(run_metis smt trivial true m2 name (these (!named_thms))) id st; ()) |
35592
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
527 |
else () |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
528 |
else |
40416
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
529 |
(Mirabelle.catch_result (metis_tag smt) false |
6461fc0f9f47
use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents:
40374
diff
changeset
|
530 |
(run_metis smt trivial false m1 name (these (!named_thms))) id st; ()) |
35592
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
531 |
in |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
532 |
change_data id (set_mini minimize); |
39337
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents:
39321
diff
changeset
|
533 |
Mirabelle.catch sh_tag (run_sledgehammer trivial args named_thms) id st; |
35592
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
534 |
if is_some (!named_thms) |
32612 | 535 |
then |
35592
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
536 |
(apply_metis Unminimized UnminimizedFT; |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
537 |
if minimize andalso not (null (these (!named_thms))) |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
538 |
then |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
539 |
(Mirabelle.catch minimize_tag (run_minimize args named_thms) id st; |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
540 |
apply_metis Minimized MinimizedFT) |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
541 |
else ()) |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
542 |
else () |
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents:
34052
diff
changeset
|
543 |
end |
32818 | 544 |
end |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
545 |
|
32511 | 546 |
fun invoke args = |
32515
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents:
32511
diff
changeset
|
547 |
let |
36373
66af0a49de39
move some sledgehammer stuff out of "atp_manager.ML"
blanchet
parents:
36294
diff
changeset
|
548 |
val _ = Sledgehammer_Isar.full_types := AList.defined (op =) args full_typesK |
32521 | 549 |
in Mirabelle.register (init, sledgehammer_action args, done) end |
32385
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
550 |
|
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff
changeset
|
551 |
end |