| author | wenzelm | 
| Tue, 10 Nov 2009 23:18:03 +0100 | |
| changeset 33604 | d4220df6fde2 | 
| parent 33316 | 6a72af4e84b8 | 
| child 34035 | 08d34921b7dd | 
| 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" | 
| 32574 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 10 | val prover_hard_timeoutK = "prover_hard_timeout" | 
| 32521 | 11 | val keepK = "keep" | 
| 12 | val full_typesK = "full_types" | |
| 32525 | 13 | val minimizeK = "minimize" | 
| 14 | val minimize_timeoutK = "minimize_timeout" | |
| 32521 | 15 | |
| 16 | fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: " | |
| 32525 | 17 | fun minimize_tag id = "#" ^ string_of_int id ^ " minimize (sledgehammer): " | 
| 32521 | 18 | fun metis_tag id = "#" ^ string_of_int id ^ " metis (sledgehammer): " | 
| 19 | ||
| 32525 | 20 | val separator = "-----" | 
| 21 | ||
| 32521 | 22 | |
| 32549 | 23 | datatype sh_data = ShData of {
 | 
| 24 | calls: int, | |
| 25 | success: int, | |
| 32585 | 26 | lemmas: int, | 
| 32818 | 27 | max_lems: int, | 
| 32549 | 28 | time_isa: int, | 
| 29 | time_atp: int, | |
| 30 | time_atp_fail: int} | |
| 31 | ||
| 32 | datatype me_data = MeData of {
 | |
| 33 | calls: int, | |
| 34 | success: int, | |
| 32676 | 35 | proofs: int, | 
| 32549 | 36 | time: int, | 
| 32550 | 37 | timeout: int, | 
| 32990 | 38 | lemmas: int * int * int, | 
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 39 | posns: Position.T list | 
| 32550 | 40 | } | 
| 32549 | 41 | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 42 | datatype min_data = MinData of {
 | 
| 32609 | 43 | succs: int, | 
| 44 | ab_ratios: int, | |
| 45 | it_ratios: int | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 46 | } | 
| 32521 | 47 | |
| 32549 | 48 | (* The first me_data component is only used if "minimize" is on. | 
| 49 | Then it records how metis behaves with un-minimized lemmas. | |
| 50 | *) | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 51 | datatype data = Data of sh_data * me_data * min_data * me_data | 
| 32549 | 52 | |
| 32818 | 53 | fun make_sh_data | 
| 54 | (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) = | |
| 55 |   ShData{calls=calls, success=success, lemmas=lemmas, max_lems=max_lems,
 | |
| 56 | time_isa=time_isa, time_atp=time_atp, time_atp_fail=time_atp_fail} | |
| 32521 | 57 | |
| 32609 | 58 | fun make_min_data (succs, ab_ratios, it_ratios) = | 
| 59 |   MinData{succs=succs, ab_ratios=ab_ratios, it_ratios=it_ratios}
 | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 60 | |
| 32990 | 61 | fun make_me_data (calls,success,proofs,time,timeout,lemmas,posns) = | 
| 32818 | 62 |   MeData{calls=calls, success=success, proofs=proofs, time=time,
 | 
| 32990 | 63 | timeout=timeout, lemmas=lemmas, posns=posns} | 
| 32549 | 64 | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 65 | val empty_data = | 
| 32818 | 66 | Data(make_sh_data (0, 0, 0, 0, 0, 0, 0), | 
| 32990 | 67 | make_me_data(0, 0, 0, 0, 0, (0,0,0), []), | 
| 32609 | 68 |        MinData{succs=0, ab_ratios=0, it_ratios=0},
 | 
| 32990 | 69 | make_me_data(0, 0, 0, 0, 0, (0,0,0), [])) | 
| 32521 | 70 | |
| 32549 | 71 | fun map_sh_data f | 
| 32818 | 72 |     (Data(ShData{calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail},
 | 
| 73 | meda0, minda, meda)) = | |
| 74 | Data (make_sh_data (f (calls,success,lemmas,max_lems, | |
| 75 | time_isa,time_atp,time_atp_fail)), | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 76 | meda0, minda, meda) | 
| 32521 | 77 | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 78 | fun map_min_data f | 
| 32609 | 79 |   (Data(shda, meda0, MinData{succs,ab_ratios,it_ratios}, meda)) =
 | 
| 80 | Data(shda, meda0, make_min_data(f(succs,ab_ratios,it_ratios)), meda) | |
| 32521 | 81 | |
| 32990 | 82 | fun map_me_data0 f (Data (shda, MeData{calls,success,proofs,time,timeout,lemmas,posns}, minda, meda)) =
 | 
| 83 | Data(shda, make_me_data(f (calls,success,proofs,time,timeout,lemmas,posns)), minda, meda) | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 84 | |
| 32990 | 85 | fun map_me_data f (Data (shda, meda0, minda, MeData{calls,success,proofs,time,timeout,lemmas,posns})) =
 | 
| 86 | Data(shda, meda0, minda, make_me_data(f (calls,success,proofs,time,timeout,lemmas,posns))) | |
| 87 | ||
| 88 | fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n)); | |
| 32521 | 89 | |
| 32818 | 90 | val inc_sh_calls = map_sh_data | 
| 91 | (fn (calls, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail) | |
| 92 | => (calls + 1, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)) | |
| 32549 | 93 | |
| 32818 | 94 | val inc_sh_success = map_sh_data | 
| 95 | (fn (calls, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail) | |
| 96 | => (calls, success + 1, lemmas,max_lems, time_isa, time_atp, time_atp_fail)) | |
| 32585 | 97 | |
| 32818 | 98 | fun inc_sh_lemmas n = map_sh_data | 
| 99 | (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) | |
| 100 | => (calls,success,lemmas+n,max_lems,time_isa,time_atp,time_atp_fail)) | |
| 32521 | 101 | |
| 32818 | 102 | fun inc_sh_max_lems n = map_sh_data | 
| 103 | (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) | |
| 104 | => (calls,success,lemmas,Int.max(max_lems,n),time_isa,time_atp,time_atp_fail)) | |
| 32549 | 105 | |
| 32818 | 106 | fun inc_sh_time_isa t = map_sh_data | 
| 107 | (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) | |
| 108 | => (calls,success,lemmas,max_lems,time_isa + t,time_atp,time_atp_fail)) | |
| 109 | ||
| 110 | fun inc_sh_time_atp t = map_sh_data | |
| 111 | (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) | |
| 112 | => (calls,success,lemmas,max_lems,time_isa,time_atp + t,time_atp_fail)) | |
| 32521 | 113 | |
| 32818 | 114 | fun inc_sh_time_atp_fail t = map_sh_data | 
| 115 | (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) | |
| 116 | => (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail + t)) | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 117 | |
| 32818 | 118 | val inc_min_succs = map_min_data | 
| 119 | (fn (succs,ab_ratios,it_ratios) => (succs+1, ab_ratios, it_ratios)) | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 120 | |
| 32818 | 121 | fun inc_min_ab_ratios r = map_min_data | 
| 122 | (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios+r, it_ratios)) | |
| 32609 | 123 | |
| 32818 | 124 | fun inc_min_it_ratios r = map_min_data | 
| 125 | (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios, it_ratios+r)) | |
| 32549 | 126 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 127 | val inc_metis_calls = map_me_data | 
| 32990 | 128 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 129 | => (calls + 1, success, proofs, time, timeout, lemmas,posns)) | |
| 32533 | 130 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 131 | val inc_metis_success = map_me_data | 
| 32990 | 132 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 133 | => (calls, success + 1, proofs, time, timeout, lemmas,posns)) | |
| 32676 | 134 | |
| 135 | val inc_metis_proofs = map_me_data | |
| 32990 | 136 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 137 | => (calls, success, proofs + 1, time, timeout, lemmas,posns)) | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 138 | |
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 139 | fun inc_metis_time t = map_me_data | 
| 32990 | 140 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 141 | => (calls, success, proofs, time + t, timeout, lemmas,posns)) | |
| 32536 | 142 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 143 | val inc_metis_timeout = map_me_data | 
| 32990 | 144 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 145 | => (calls, success, proofs, time, timeout + 1, lemmas,posns)) | |
| 32549 | 146 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 147 | fun inc_metis_lemmas n = map_me_data | 
| 32990 | 148 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 149 | => (calls, success, proofs, time, timeout, inc_max n lemmas, posns)) | |
| 32550 | 150 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 151 | fun inc_metis_posns pos = map_me_data | 
| 32990 | 152 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 153 | => (calls, success, proofs, time, timeout, lemmas, pos::posns)) | |
| 32521 | 154 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 155 | val inc_metis_calls0 = map_me_data0 | 
| 32990 | 156 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 157 | => (calls + 1, success, proofs, time, timeout, lemmas,posns)) | |
| 32549 | 158 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 159 | val inc_metis_success0 = map_me_data0 | 
| 32990 | 160 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 161 | => (calls, success + 1, proofs, time, timeout, lemmas,posns)) | |
| 32676 | 162 | |
| 163 | val inc_metis_proofs0 = map_me_data0 | |
| 32990 | 164 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 165 | => (calls, success, proofs + 1, time, timeout, lemmas,posns)) | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 166 | |
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 167 | fun inc_metis_time0 t = map_me_data0 | 
| 32990 | 168 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 169 | => (calls, success, proofs, time + t, timeout, lemmas,posns)) | |
| 32521 | 170 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 171 | val inc_metis_timeout0 = map_me_data0 | 
| 32990 | 172 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 173 | => (calls, success, proofs, time, timeout + 1, lemmas,posns)) | |
| 32549 | 174 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 175 | fun inc_metis_lemmas0 n = map_me_data0 | 
| 32990 | 176 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 177 | => (calls, success, proofs, time, timeout, inc_max n lemmas, posns)) | |
| 32521 | 178 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 179 | fun inc_metis_posns0 pos = map_me_data0 | 
| 32990 | 180 | (fn (calls,success,proofs,time,timeout,lemmas,posns) | 
| 181 | => (calls, success, proofs, time, timeout, lemmas, pos::posns)) | |
| 32521 | 182 | |
| 183 | local | |
| 184 | ||
| 185 | val str = string_of_int | |
| 186 | val str3 = Real.fmt (StringCvt.FIX (SOME 3)) | |
| 187 | fun percentage a b = string_of_int (a * 100 div b) | |
| 188 | fun time t = Real.fromInt t / 1000.0 | |
| 189 | fun avg_time t n = | |
| 190 | if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0 | |
| 191 | ||
| 32818 | 192 | fun log_sh_data log calls success lemmas max_lems time_isa time_atp time_atp_fail = | 
| 193 |  (log ("Total number of sledgehammer calls: " ^ str calls);
 | |
| 194 |   log ("Number of successful sledgehammer calls: " ^ str success);
 | |
| 195 |   log ("Number of sledgehammer lemmas: " ^ str lemmas);
 | |
| 196 |   log ("Max number of sledgehammer lemmas: " ^ str max_lems);
 | |
| 197 |   log ("Success rate: " ^ percentage success calls ^ "%");
 | |
| 198 |   log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
 | |
| 199 |   log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_atp));
 | |
| 200 |   log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_atp_fail));
 | |
| 32536 | 201 |   log ("Average time for sledgehammer calls (Isabelle): " ^
 | 
| 32818 | 202 | str3 (avg_time time_isa calls)); | 
| 32533 | 203 |   log ("Average time for successful sledgehammer calls (ATP): " ^
 | 
| 32818 | 204 | str3 (avg_time time_atp success)); | 
| 32536 | 205 |   log ("Average time for failed sledgehammer calls (ATP): " ^
 | 
| 32818 | 206 | str3 (avg_time time_atp_fail (calls - success))) | 
| 32533 | 207 | ) | 
| 32521 | 208 | |
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 209 | |
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 210 | fun str_of_pos pos = | 
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 211 | let val str0 = string_of_int o the_default 0 | 
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 212 | in str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) end | 
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 213 | |
| 32676 | 214 | fun log_metis_data log tag sh_calls sh_success metis_calls metis_success metis_proofs metis_time | 
| 32990 | 215 | metis_timeout (lemmas, lems_sos, lems_max) metis_posns = | 
| 32549 | 216 |  (log ("Total number of " ^ tag ^ "metis calls: " ^ str metis_calls);
 | 
| 32740 | 217 |   log ("Number of successful " ^ tag ^ "metis calls: " ^ str metis_success ^
 | 
| 218 | " (proof: " ^ str metis_proofs ^ ")"); | |
| 32549 | 219 |   log ("Number of " ^ tag ^ "metis timeouts: " ^ str metis_timeout);
 | 
| 32533 | 220 |   log ("Success rate: " ^ percentage metis_success sh_calls ^ "%");
 | 
| 32990 | 221 |   log ("Number of successful " ^ tag ^ "metis lemmas: " ^ str lemmas);
 | 
| 222 |   log ("SOS of successful " ^ tag ^ "metis lemmas: " ^ str lems_sos);
 | |
| 223 |   log ("Max number of successful " ^ tag ^ "metis lemmas: " ^ str lems_max);
 | |
| 224 |   log ("Total time for successful " ^ tag ^ "metis calls: " ^ str3 (time metis_time));
 | |
| 32521 | 225 |   log ("Average time for successful metis calls: " ^
 | 
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 226 | str3 (avg_time metis_time metis_success)); | 
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 227 | if tag="" | 
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 228 |   then log ("Proved: " ^ space_implode " " (map str_of_pos metis_posns))
 | 
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 229 | else () | 
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 230 | ) | 
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 231 | |
| 32609 | 232 | fun log_min_data log succs ab_ratios it_ratios = | 
| 233 |   (log ("Number of successful minimizations: " ^ string_of_int succs);
 | |
| 234 |    log ("After/before ratios: " ^ string_of_int ab_ratios);
 | |
| 235 |    log ("Iterations ratios: " ^ string_of_int it_ratios)
 | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 236 | ) | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 237 | |
| 32521 | 238 | in | 
| 239 | ||
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 240 | fun log_data id log (Data | 
| 32818 | 241 |    (ShData{calls=sh_calls, lemmas=sh_lemmas,  max_lems=sh_max_lems, success=sh_success,
 | 
| 32585 | 242 | time_isa=sh_time_isa,time_atp=sh_time_atp,time_atp_fail=sh_time_atp_fail}, | 
| 32676 | 243 |     MeData{calls=metis_calls0, proofs=metis_proofs0,
 | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 244 | success=metis_success0, time=metis_time0, timeout=metis_timeout0, | 
| 32990 | 245 | lemmas=metis_lemmas0, posns=metis_posns0}, | 
| 32609 | 246 |     MinData{succs=min_succs, ab_ratios=ab_ratios, it_ratios=it_ratios},
 | 
| 32676 | 247 |     MeData{calls=metis_calls, proofs=metis_proofs,
 | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 248 | success=metis_success, time=metis_time, timeout=metis_timeout, | 
| 32990 | 249 | lemmas=metis_lemmas, posns=metis_posns})) = | 
| 32521 | 250 | if sh_calls > 0 | 
| 251 | then | |
| 252 |    (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
 | |
| 32818 | 253 | log_sh_data log sh_calls sh_success sh_lemmas sh_max_lems sh_time_isa sh_time_atp sh_time_atp_fail; | 
| 32521 | 254 | log ""; | 
| 32549 | 255 | if metis_calls > 0 then log_metis_data log "" sh_calls sh_success metis_calls | 
| 32990 | 256 | metis_success metis_proofs metis_time metis_timeout metis_lemmas metis_posns else (); | 
| 32549 | 257 | log ""; | 
| 258 | if metis_calls0 > 0 | |
| 32609 | 259 | then (log_min_data log min_succs ab_ratios it_ratios; log ""; | 
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 260 | log_metis_data log "unminimized " sh_calls sh_success metis_calls0 | 
| 32990 | 261 | metis_success0 metis_proofs0 metis_time0 metis_timeout0 metis_lemmas0 metis_posns0) | 
| 32549 | 262 | else () | 
| 263 | ) | |
| 32521 | 264 | else () | 
| 265 | ||
| 266 | end | |
| 267 | ||
| 268 | ||
| 269 | (* Warning: we implicitly assume single-threaded execution here! *) | |
| 32740 | 270 | val data = Unsynchronized.ref ([] : (int * data) list) | 
| 32521 | 271 | |
| 32740 | 272 | fun init id thy = (Unsynchronized.change data (cons (id, empty_data)); thy) | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 273 | fun done id ({log, ...}: Mirabelle.done_args) =
 | 
| 32521 | 274 | AList.lookup (op =) (!data) id | 
| 275 | |> Option.map (log_data id log) | |
| 276 | |> K () | |
| 277 | ||
| 32740 | 278 | fun change_data id f = (Unsynchronized.change data (AList.map_entry (op =) id f); ()) | 
| 32521 | 279 | |
| 280 | ||
| 32525 | 281 | fun get_atp thy args = | 
| 33016 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 282 | let | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 283 | fun default_atp_name () = hd (ATP_Manager.get_atps ()) | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 284 | handle Empty => error "No ATP available." | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 285 | fun get_prover name = | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 286 | (case ATP_Manager.get_prover thy name of | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 287 | SOME prover => (name, prover) | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 288 |       | NONE => error ("Bad ATP: " ^ quote name))
 | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 289 | in | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 290 | (case AList.lookup (op =) args proverK of | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 291 | SOME name => get_prover name | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 292 | | NONE => get_prover (default_atp_name ())) | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 293 | end | 
| 32525 | 294 | |
| 32521 | 295 | local | 
| 296 | ||
| 32536 | 297 | datatype sh_result = | 
| 298 | SH_OK of int * int * string list | | |
| 299 | SH_FAIL of int * int | | |
| 300 | SH_ERROR | |
| 301 | ||
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32819diff
changeset | 302 | fun run_sh prover hard_timeout timeout dir st = | 
| 32521 | 303 | let | 
| 33292 | 304 |     val {context = ctxt, facts, goal} = Proof.raw_goal st   (* FIXME ?? *)
 | 
| 33247 | 305 | val change_dir = (fn SOME d => Config.put ATP_Wrapper.destdir d | _ => I) | 
| 306 | val ctxt' = | |
| 307 | ctxt | |
| 308 | |> change_dir dir | |
| 309 | |> Config.put ATP_Wrapper.measure_runtime true | |
| 33292 | 310 | val problem = ATP_Wrapper.problem_of_goal (! ATP_Manager.full_types) 1 (ctxt', (facts, goal)); | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32819diff
changeset | 311 | |
| 32574 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 312 | val time_limit = | 
| 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 313 | (case hard_timeout of | 
| 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 314 | NONE => I | 
| 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 315 | | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs)) | 
| 32985 | 316 |     val ({success, message, theorem_names, runtime=time_atp, ...}: ATP_Wrapper.prover_result,
 | 
| 317 | time_isa) = time_limit (Mirabelle.cpu_time (fn t => prover t problem)) timeout | |
| 32521 | 318 | in | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32819diff
changeset | 319 | if success then (message, SH_OK (time_isa, time_atp, theorem_names)) | 
| 32536 | 320 | else (message, SH_FAIL(time_isa, time_atp)) | 
| 32521 | 321 | end | 
| 33316 | 322 |   handle Res_HOL_Clause.TOO_TRIVIAL => ("trivial", SH_OK (0, 0, []))
 | 
| 32536 | 323 |        | 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: 
32571diff
changeset | 324 |        | TimeLimit.TimeOut => ("timeout", SH_ERROR)
 | 
| 32521 | 325 | |
| 32454 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 326 | 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: 
32452diff
changeset | 327 | let | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 328 | val lex = OuterKeyword.get_lexicons | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 329 | 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: 
32452diff
changeset | 330 | in | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 331 | Source.of_string name | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 332 |     |> Symbol.source {do_recover=false}
 | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 333 |     |> OuterLex.source {do_recover=SOME false} lex Position.start
 | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 334 | |> OuterLex.source_proper | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 335 | |> Source.source OuterLex.stopper (SpecParse.xthms1 >> get) NONE | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 336 | |> Source.exhaust | 
| 
a1a5589207ad
Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
 boehmes parents: 
32452diff
changeset | 337 | end | 
| 32452 
d84edd022efe
apply metis with found theorems in case sledgehammer was successful
 boehmes parents: 
32434diff
changeset | 338 | |
| 32498 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32496diff
changeset | 339 | in | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32496diff
changeset | 340 | |
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 341 | fun run_sledgehammer 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 | 342 | let | 
| 32536 | 343 | val _ = change_data id inc_sh_calls | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32819diff
changeset | 344 | val (prover_name, prover) = get_atp (Proof.theory_of st) args | 
| 32525 | 345 | val dir = AList.lookup (op =) args keepK | 
| 32541 | 346 | val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30) | 
| 32574 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 347 | val hard_timeout = AList.lookup (op =) args prover_hard_timeoutK | 
| 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 348 | |> Option.map (fst o read_int o explode) | 
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32819diff
changeset | 349 | val (msg, result) = run_sh prover hard_timeout timeout dir st | 
| 32525 | 350 | in | 
| 32536 | 351 | case result of | 
| 352 | SH_OK (time_isa, time_atp, names) => | |
| 32818 | 353 | let fun get_thms name = (name, thms_of_name (Proof.context_of st) name) | 
| 32525 | 354 | in | 
| 32818 | 355 | change_data id inc_sh_success; | 
| 356 | change_data id (inc_sh_lemmas (length names)); | |
| 357 | change_data id (inc_sh_max_lems (length names)); | |
| 358 | change_data id (inc_sh_time_isa time_isa); | |
| 359 | change_data id (inc_sh_time_atp time_atp); | |
| 360 | named_thms := SOME (map get_thms names); | |
| 32536 | 361 |           log (sh_tag id ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
 | 
| 362 | string_of_int time_atp ^ ") [" ^ prover_name ^ "]:\n" ^ msg) | |
| 32525 | 363 | end | 
| 32536 | 364 | | SH_FAIL (time_isa, time_atp) => | 
| 365 | let | |
| 366 | val _ = change_data id (inc_sh_time_isa time_isa) | |
| 367 | val _ = change_data id (inc_sh_time_atp_fail time_atp) | |
| 368 | in log (sh_tag id ^ "failed: " ^ msg) end | |
| 369 | | SH_ERROR => log (sh_tag id ^ "failed: " ^ msg) | |
| 32525 | 370 | end | 
| 371 | ||
| 372 | end | |
| 373 | ||
| 32521 | 374 | |
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 375 | fun run_minimize args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
 | 
| 32525 | 376 | let | 
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 377 | val n0 = length (these (!named_thms)) | 
| 32525 | 378 | val (prover_name, prover) = get_atp (Proof.theory_of st) args | 
| 32936 | 379 | val minimize = ATP_Minimal.minimalize prover prover_name | 
| 32525 | 380 | val timeout = | 
| 381 | AList.lookup (op =) args minimize_timeoutK | |
| 382 | |> Option.map (fst o read_int o explode) | |
| 383 | |> the_default 5 | |
| 384 | val _ = log separator | |
| 385 | in | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 386 | case minimize timeout st (these (!named_thms)) of | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 387 | (SOME (named_thms',its), msg) => | 
| 32609 | 388 | (change_data id inc_min_succs; | 
| 389 | change_data id (inc_min_ab_ratios ((100 * length named_thms') div n0)); | |
| 390 | change_data id (inc_min_it_ratios ((100*its) div n0)); | |
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 391 | if length named_thms' = n0 | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 392 | then log (minimize_tag id ^ "already minimal") | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 393 | else (named_thms := SOME named_thms'; | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 394 | log (minimize_tag id ^ "succeeded:\n" ^ msg)) | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 395 | ) | 
| 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 396 | | (NONE, msg) => log (minimize_tag id ^ "failed: " ^ msg) | 
| 32525 | 397 | end | 
| 398 | ||
| 399 | ||
| 32676 | 400 | fun run_metis (inc_metis_calls, inc_metis_success, inc_metis_proofs, inc_metis_time, inc_metis_timeout, | 
| 32990 | 401 | inc_metis_lemmas, inc_metis_posns) args name named_thms id | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 402 |     ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
 | 
| 32525 | 403 | let | 
| 32521 | 404 | fun metis thms ctxt = MetisTools.metis_tac ctxt thms | 
| 405 | fun apply_metis thms = Mirabelle.can_apply timeout (metis thms) st | |
| 406 | ||
| 407 |     fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
 | |
| 408 | | with_time (true, t) = (change_data id inc_metis_success; | |
| 32585 | 409 | change_data id (inc_metis_lemmas (length named_thms)); | 
| 32521 | 410 | change_data id (inc_metis_time t); | 
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 411 | change_data id (inc_metis_posns pos); | 
| 32676 | 412 | if name = "proof" then change_data id inc_metis_proofs else (); | 
| 32521 | 413 |           "succeeded (" ^ string_of_int t ^ ")")
 | 
| 414 | fun timed_metis thms = with_time (Mirabelle.cpu_time apply_metis thms) | |
| 415 | handle TimeLimit.TimeOut => (change_data id inc_metis_timeout; "timeout") | |
| 416 | | ERROR msg => "error: " ^ msg | |
| 417 | ||
| 32525 | 418 | val _ = log separator | 
| 32521 | 419 | val _ = change_data id inc_metis_calls | 
| 420 | in | |
| 32525 | 421 | maps snd named_thms | 
| 32521 | 422 | |> timed_metis | 
| 423 | |> log o prefix (metis_tag id) | |
| 424 | end | |
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: diff
changeset | 425 | |
| 32676 | 426 | fun sledgehammer_action args id (st as {log, pre, name, ...}: Mirabelle.run_args) =
 | 
| 33292 | 427 | let val goal = Thm.major_prem_of (#goal (Proof.raw_goal pre)) (* FIXME ?? *) in | 
| 32818 | 428 | if can Logic.dest_conjunction goal orelse can Logic.dest_equals goal | 
| 32607 | 429 | then () else | 
| 32515 
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
 boehmes parents: 
32511diff
changeset | 430 | let | 
| 32676 | 431 | val metis_fns = (inc_metis_calls, inc_metis_success, inc_metis_proofs, inc_metis_time, | 
| 32990 | 432 | inc_metis_timeout, inc_metis_lemmas, inc_metis_posns) | 
| 32676 | 433 | val metis0_fns = (inc_metis_calls0, inc_metis_success0, inc_metis_proofs0, inc_metis_time0, | 
| 32990 | 434 | inc_metis_timeout0, inc_metis_lemmas0, inc_metis_posns0) | 
| 32740 | 435 | val named_thms = Unsynchronized.ref (NONE : (string * thm list) list option) | 
| 32612 | 436 | val minimize = AList.defined (op =) args minimizeK | 
| 437 | in | |
| 438 | Mirabelle.catch sh_tag (run_sledgehammer args named_thms) id st; | |
| 439 | if is_some (!named_thms) | |
| 440 | then | |
| 441 | (if minimize | |
| 32676 | 442 | then Mirabelle.catch metis_tag (run_metis metis0_fns args name (these (!named_thms))) id st | 
| 32612 | 443 | else (); | 
| 444 | if minimize andalso not(null(these(!named_thms))) | |
| 445 | then Mirabelle.catch minimize_tag (run_minimize args named_thms) id st | |
| 446 | else (); | |
| 32676 | 447 | Mirabelle.catch metis_tag (run_metis metis_fns args name (these (!named_thms))) id st) | 
| 32612 | 448 | else () | 
| 449 | end | |
| 32818 | 450 | end | 
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: diff
changeset | 451 | |
| 32511 | 452 | fun invoke args = | 
| 32515 
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
 boehmes parents: 
32511diff
changeset | 453 | let | 
| 32937 
34f66c9dd8a2
structure ATP_Manager: eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
 wenzelm parents: 
32936diff
changeset | 454 | val _ = ATP_Manager.full_types := AList.defined (op =) args full_typesK | 
| 32521 | 455 | 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 | 456 | |
| 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: diff
changeset | 457 | end |