| author | wenzelm | 
| Mon, 19 Jan 2015 16:31:04 +0100 | |
| changeset 59406 | 283aa6225d98 | 
| parent 59083 | 88b0b1f28adc | 
| child 59582 | 0fbed69ff081 | 
| permissions | -rw-r--r-- | 
| 47847 | 1 | (* Title: HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML | 
| 32564 | 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 | |
| 47480 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 8 | (*To facilitate synching the description of Mirabelle Sledgehammer parameters | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 9 | (in ../lib/Tools/mirabelle) with the parameters actually used by this | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 10 | interface, the former extracts PARAMETER and DESCRIPTION from code below which | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 11 | has this pattern (provided it appears in a single line): | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 12 | val .*K = "PARAMETER" (*DESCRIPTION*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 13 | *) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 14 | (*NOTE: descriptions mention parameters (particularly NAME) without a defined range.*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 15 | val proverK = "prover" (*=NAME: name of the external prover to call*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 16 | val prover_timeoutK = "prover_timeout" (*=TIME: timeout for invoked ATP (seconds of process time)*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 17 | val keepK = "keep" (*=PATH: path where to keep temporary files created by sledgehammer*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 18 | |
| 55285 | 19 | val proof_methodK = "proof_method" (*=NAME: how to reconstruct proofs (ie. using metis/smt)*) | 
| 47480 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 20 | |
| 48293 | 21 | val max_factsK = "max_facts" (*=NUM: max. relevant clauses to use*) | 
| 47480 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 22 | val max_relevantK = "max_relevant" (*=NUM: max. relevant clauses to use*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 23 | val max_callsK = "max_calls" (*=NUM: max. no. of calls to sledgehammer*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 24 | val preplay_timeoutK = "preplay_timeout" (*=TIME: timeout for finding reconstructed proof*) | 
| 56411 | 25 | val isar_proofsK = "isar_proofs" (*: enable Isar proof generation*) | 
| 58473 | 26 | val smt_proofsK = "smt_proofs" (*: enable SMT proof generation*) | 
| 57782 | 27 | val minimizeK = "minimize" (*: instruct sledgehammer to run its minimizer*) | 
| 47480 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 28 | |
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 29 | val check_trivialK = "check_trivial" (*: check if goals are "trivial" (false by default)*) | 
| 50334 | 30 | val fact_filterK = "fact_filter" (*=STRING: fact filter*) | 
| 47480 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 31 | val type_encK = "type_enc" (*=STRING: type encoding scheme*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 32 | val lam_transK = "lam_trans" (*=STRING: lambda translation scheme*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 33 | val strictK = "strict" (*=BOOL: run in strict mode*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 34 | val sliceK = "slice" (*=BOOL: allow sledgehammer-level strategy-scheduling*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 35 | val uncurried_aliasesK = "uncurried_aliases" (*=SMART_BOOL: use fresh function names to alias curried applications*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 36 | val e_selection_heuristicK = "e_selection_heuristic" (*: FIXME*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 37 | val term_orderK = "term_order" (*: FIXME*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 38 | val force_sosK = "force_sos" (*: use SOS*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 39 | val max_new_mono_instancesK = "max_new_mono_instances" (*=NUM: max. new monomorphic instances*) | 
| 
24d8c9e9dae4
Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
 sultana parents: 
47477diff
changeset | 40 | val max_mono_itersK = "max_mono_iters" (*=NUM: max. iterations of monomorphiser*) | 
| 44448 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 41 | |
| 32521 | 42 | fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: " | 
| 55285 | 43 | fun proof_method_tag meth id = "#" ^ string_of_int id ^ " " ^ (!meth) ^ " (sledgehammer): " | 
| 32521 | 44 | |
| 32525 | 45 | val separator = "-----" | 
| 46 | ||
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 47 | (*FIXME sensible to have Mirabelle-level Sledgehammer defaults?*) | 
| 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 48 | (*defaults used in this Mirabelle action*) | 
| 57751 
f453bbc289fa
modernized Mirabelle (a bit) and made it compile
 blanchet parents: 
57154diff
changeset | 49 | val preplay_timeout_default = "1" | 
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 50 | val lam_trans_default = "smart" | 
| 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 51 | val uncurried_aliases_default = "smart" | 
| 50334 | 52 | val fact_filter_default = "smart" | 
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 53 | val type_enc_default = "smart" | 
| 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 54 | val strict_default = "false" | 
| 48293 | 55 | val max_facts_default = "smart" | 
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 56 | val slice_default = "true" | 
| 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 57 | val max_calls_default = "10000000" | 
| 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 58 | val trivial_default = "false" | 
| 46826 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 59 | |
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 60 | (*If a key is present in args then augment a list with its pair*) | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 61 | (*This is used to avoid fixing default values at the Mirabelle level, and | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 62 | instead use the default values of the tool (Sledgehammer in this case).*) | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 63 | fun available_parameter args key label list = | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 64 | let | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 65 | val value = AList.lookup (op =) args key | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 66 | in if is_some value then (label, the value) :: list else list end | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 67 | |
| 32549 | 68 | datatype sh_data = ShData of {
 | 
| 69 | calls: int, | |
| 70 | success: int, | |
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 71 | nontriv_calls: int, | 
| 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 72 | nontriv_success: int, | 
| 32585 | 73 | lemmas: int, | 
| 32818 | 74 | max_lems: int, | 
| 32549 | 75 | time_isa: int, | 
| 40062 | 76 | time_prover: int, | 
| 77 | time_prover_fail: int} | |
| 32549 | 78 | |
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 79 | datatype re_data = ReData of {
 | 
| 32549 | 80 | calls: int, | 
| 81 | success: int, | |
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 82 | nontriv_calls: int, | 
| 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 83 | nontriv_success: int, | 
| 32676 | 84 | proofs: int, | 
| 32549 | 85 | time: int, | 
| 32550 | 86 | timeout: int, | 
| 32990 | 87 | lemmas: int * int * int, | 
| 39341 
d2b981a0429a
indicate triviality in the list of proved things
 blanchet parents: 
39340diff
changeset | 88 | posns: (Position.T * bool) list | 
| 32550 | 89 | } | 
| 32549 | 90 | |
| 32818 | 91 | fun make_sh_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 92 | (calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa, | 
| 40062 | 93 | time_prover,time_prover_fail) = | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 94 |   ShData{calls=calls, success=success, nontriv_calls=nontriv_calls,
 | 
| 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 95 | nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems, | 
| 40062 | 96 | time_isa=time_isa, time_prover=time_prover, | 
| 97 | time_prover_fail=time_prover_fail} | |
| 32521 | 98 | |
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 99 | fun make_re_data (calls,success,nontriv_calls,nontriv_success,proofs,time, | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 100 | timeout,lemmas,posns) = | 
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 101 |   ReData{calls=calls, success=success, nontriv_calls=nontriv_calls,
 | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 102 | nontriv_success=nontriv_success, proofs=proofs, time=time, | 
| 32990 | 103 | timeout=timeout, lemmas=lemmas, posns=posns} | 
| 32549 | 104 | |
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 105 | val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0, 0, 0) | 
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 106 | val empty_re_data = make_re_data (0, 0, 0, 0, 0, 0, 0, (0,0,0), []) | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 107 | |
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 108 | 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: 
39321diff
changeset | 109 | lemmas, max_lems, time_isa, | 
| 40062 | 110 | time_prover, time_prover_fail}) = (calls, success, nontriv_calls, | 
| 111 | 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: 
33316diff
changeset | 112 | |
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 113 | fun tuple_of_re_data (ReData {calls, success, nontriv_calls, nontriv_success,
 | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 114 | proofs, time, timeout, lemmas, posns}) = (calls, success, nontriv_calls, | 
| 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 115 | nontriv_success, proofs, time, timeout, lemmas, posns) | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 116 | |
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 117 | datatype data = Data of {
 | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 118 | sh: sh_data, | 
| 57753 | 119 | re_u: re_data (* proof method with unminimized set of lemmas *) | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 120 | } | 
| 32521 | 121 | |
| 57753 | 122 | fun make_data (sh, re_u) = Data {sh=sh, re_u=re_u}
 | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 123 | |
| 57753 | 124 | val empty_data = make_data (empty_sh_data, empty_re_data) | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 125 | |
| 57753 | 126 | fun map_sh_data f (Data {sh, re_u}) =
 | 
| 57752 | 127 | let val sh' = make_sh_data (f (tuple_of_sh_data sh)) | 
| 57753 | 128 | in make_data (sh', re_u) end | 
| 32521 | 129 | |
| 57753 | 130 | fun map_re_data f (Data {sh, re_u}) =
 | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 131 | let | 
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 132 | val f' = make_re_data o f o tuple_of_re_data | 
| 57753 | 133 | val re_u' = f' re_u | 
| 134 | in make_data (sh, re_u') end | |
| 32990 | 135 | |
| 136 | fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n)); | |
| 32521 | 137 | |
| 32818 | 138 | val inc_sh_calls = map_sh_data | 
| 40062 | 139 | (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail) | 
| 140 | => (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail)) | |
| 32549 | 141 | |
| 32818 | 142 | val inc_sh_success = map_sh_data | 
| 40062 | 143 | (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail) | 
| 144 | => (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: 
39321diff
changeset | 145 | |
| 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 146 | val inc_sh_nontriv_calls = map_sh_data | 
| 40062 | 147 | (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail) | 
| 148 | => (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: 
39321diff
changeset | 149 | |
| 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 150 | val inc_sh_nontriv_success = map_sh_data | 
| 40062 | 151 | (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail) | 
| 152 | => (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_prover, time_prover_fail)) | |
| 32585 | 153 | |
| 32818 | 154 | fun inc_sh_lemmas n = map_sh_data | 
| 40062 | 155 | (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) | 
| 156 | => (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_prover,time_prover_fail)) | |
| 32521 | 157 | |
| 32818 | 158 | fun inc_sh_max_lems n = map_sh_data | 
| 40062 | 159 | (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) | 
| 160 | => (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_prover,time_prover_fail)) | |
| 32549 | 161 | |
| 32818 | 162 | fun inc_sh_time_isa t = map_sh_data | 
| 40062 | 163 | (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) | 
| 164 | => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_prover,time_prover_fail)) | |
| 32818 | 165 | |
| 40062 | 166 | fun inc_sh_time_prover t = map_sh_data | 
| 167 | (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) | |
| 168 | => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover + t,time_prover_fail)) | |
| 32521 | 169 | |
| 40062 | 170 | fun inc_sh_time_prover_fail t = map_sh_data | 
| 171 | (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail) | |
| 172 | => (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: 
32567diff
changeset | 173 | |
| 55285 | 174 | val inc_proof_method_calls = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 175 | (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: 
39321diff
changeset | 176 | => (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns)) | 
| 32533 | 177 | |
| 55285 | 178 | val inc_proof_method_success = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 179 | (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: 
39321diff
changeset | 180 | => (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: 
39321diff
changeset | 181 | |
| 55285 | 182 | val inc_proof_method_nontriv_calls = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 183 | (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: 
39321diff
changeset | 184 | => (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: 
39321diff
changeset | 185 | |
| 55285 | 186 | val inc_proof_method_nontriv_success = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 187 | (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: 
39321diff
changeset | 188 | => (calls, success, nontriv_calls, nontriv_success + 1, proofs, time, timeout, lemmas,posns)) | 
| 32676 | 189 | |
| 55285 | 190 | val inc_proof_method_proofs = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 191 | (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: 
39321diff
changeset | 192 | => (calls, success, nontriv_calls, nontriv_success, proofs + 1, time, timeout, lemmas,posns)) | 
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 193 | |
| 57753 | 194 | fun inc_proof_method_time t = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 195 | (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns) | 
| 57753 | 196 | => (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns)) | 
| 32536 | 197 | |
| 55285 | 198 | val inc_proof_method_timeout = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 199 | (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: 
39321diff
changeset | 200 | => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout + 1, lemmas,posns)) | 
| 32549 | 201 | |
| 57753 | 202 | fun inc_proof_method_lemmas n = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 203 | (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns) | 
| 57753 | 204 | => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, inc_max n lemmas, posns)) | 
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 205 | |
| 57753 | 206 | fun inc_proof_method_posns pos = map_re_data | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 207 | (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns) | 
| 57753 | 208 | => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns)) | 
| 32521 | 209 | |
| 44090 | 210 | val str0 = string_of_int o the_default 0 | 
| 211 | ||
| 32521 | 212 | local | 
| 213 | ||
| 214 | val str = string_of_int | |
| 215 | val str3 = Real.fmt (StringCvt.FIX (SOME 3)) | |
| 216 | fun percentage a b = string_of_int (a * 100 div b) | |
| 217 | fun time t = Real.fromInt t / 1000.0 | |
| 218 | fun avg_time t n = | |
| 219 | if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0 | |
| 220 | ||
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 221 | fun log_sh_data log | 
| 40062 | 222 | (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail) = | 
| 32818 | 223 |  (log ("Total number of sledgehammer calls: " ^ str calls);
 | 
| 224 |   log ("Number of successful sledgehammer calls: " ^ str success);
 | |
| 225 |   log ("Number of sledgehammer lemmas: " ^ str lemmas);
 | |
| 226 |   log ("Max number of sledgehammer lemmas: " ^ str max_lems);
 | |
| 227 |   log ("Success rate: " ^ percentage success calls ^ "%");
 | |
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 228 |   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: 
39321diff
changeset | 229 |   log ("Number of successful nontrivial sledgehammer calls: " ^ str nontriv_success);
 | 
| 32818 | 230 |   log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
 | 
| 40062 | 231 |   log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_prover));
 | 
| 232 |   log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_prover_fail));
 | |
| 32536 | 233 |   log ("Average time for sledgehammer calls (Isabelle): " ^
 | 
| 32818 | 234 | str3 (avg_time time_isa calls)); | 
| 32533 | 235 |   log ("Average time for successful sledgehammer calls (ATP): " ^
 | 
| 40062 | 236 | str3 (avg_time time_prover success)); | 
| 32536 | 237 |   log ("Average time for failed sledgehammer calls (ATP): " ^
 | 
| 40062 | 238 | str3 (avg_time time_prover_fail (calls - success))) | 
| 32533 | 239 | ) | 
| 32521 | 240 | |
| 39341 
d2b981a0429a
indicate triviality in the list of proved things
 blanchet parents: 
39340diff
changeset | 241 | fun str_of_pos (pos, triv) = | 
| 47728 
6ee015f6ea4b
reintroduce file offsets in Mirabelle output, but make sure they are not influenced by the length of the path
 blanchet parents: 
47532diff
changeset | 242 | str0 (Position.line_of pos) ^ ":" ^ str0 (Position.offset_of pos) ^ | 
| 44090 | 243 | (if triv then "[T]" else "") | 
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 244 | |
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 245 | fun log_re_data log tag sh_calls (re_calls, re_success, re_nontriv_calls, | 
| 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 246 | re_nontriv_success, re_proofs, re_time, re_timeout, | 
| 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 247 | (lemmas, lems_sos, lems_max), re_posns) = | 
| 55285 | 248 |  (log ("Total number of " ^ tag ^ "proof method calls: " ^ str re_calls);
 | 
| 249 |   log ("Number of successful " ^ tag ^ "proof method calls: " ^ str re_success ^
 | |
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 250 | " (proof: " ^ str re_proofs ^ ")"); | 
| 55285 | 251 |   log ("Number of " ^ tag ^ "proof method timeouts: " ^ str re_timeout);
 | 
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 252 |   log ("Success rate: " ^ percentage re_success sh_calls ^ "%");
 | 
| 55285 | 253 |   log ("Total number of nontrivial " ^ tag ^ "proof method calls: " ^ str re_nontriv_calls);
 | 
| 254 |   log ("Number of successful nontrivial " ^ tag ^ "proof method calls: " ^ str re_nontriv_success ^
 | |
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 255 | " (proof: " ^ str re_proofs ^ ")"); | 
| 55285 | 256 |   log ("Number of successful " ^ tag ^ "proof method lemmas: " ^ str lemmas);
 | 
| 257 |   log ("SOS of successful " ^ tag ^ "proof method lemmas: " ^ str lems_sos);
 | |
| 258 |   log ("Max number of successful " ^ tag ^ "proof method lemmas: " ^ str lems_max);
 | |
| 259 |   log ("Total time for successful " ^ tag ^ "proof method calls: " ^ str3 (time re_time));
 | |
| 260 |   log ("Average time for successful " ^ tag ^ "proof method calls: " ^
 | |
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 261 | str3 (avg_time re_time re_success)); | 
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 262 | if tag="" | 
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 263 |   then log ("Proved: " ^ space_implode " " (map str_of_pos re_posns))
 | 
| 32551 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 264 | else () | 
| 
421323205efd
position information is now passed to all actions;
 nipkow parents: 
32550diff
changeset | 265 | ) | 
| 32571 
d4bb776874b8
count number of iterations required for minimization (and fixed bug: minimization was always called)
 nipkow parents: 
32567diff
changeset | 266 | |
| 32521 | 267 | in | 
| 268 | ||
| 57753 | 269 | fun log_data id log (Data {sh, re_u}) =
 | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 270 | let | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 271 |     val ShData {calls=sh_calls, ...} = sh
 | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 272 | |
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 273 |     fun app_if (ReData {calls, ...}) f = if calls > 0 then f () else ()
 | 
| 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 274 | fun log_re tag m = | 
| 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 275 | log_re_data log tag sh_calls (tuple_of_re_data m) | 
| 57753 | 276 | fun log_proof_method (tag1, m1) = app_if m1 (fn () => (log_re tag1 m1; log "")) | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 277 | in | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 278 | if sh_calls > 0 | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 279 | then | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 280 |      (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
 | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 281 | log_sh_data log (tuple_of_sh_data sh); | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 282 | log ""; | 
| 57753 | 283 |       log_proof_method ("", re_u))
 | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 284 | else () | 
| 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 285 | end | 
| 32521 | 286 | |
| 287 | end | |
| 288 | ||
| 55211 | 289 | (* Warning: we implicitly assume single-threaded execution here *) | 
| 32740 | 290 | val data = Unsynchronized.ref ([] : (int * data) list) | 
| 32521 | 291 | |
| 32740 | 292 | 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 | 293 | fun done id ({log, ...}: Mirabelle.done_args) =
 | 
| 32521 | 294 | AList.lookup (op =) (!data) id | 
| 295 | |> Option.map (log_data id log) | |
| 296 | |> K () | |
| 297 | ||
| 32740 | 298 | fun change_data id f = (Unsynchronized.change data (AList.map_entry (op =) id f); ()) | 
| 32521 | 299 | |
| 55205 | 300 | fun get_prover_name thy args = | 
| 33016 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 301 | let | 
| 40062 | 302 | fun default_prover_name () = | 
| 55205 | 303 | hd (#provers (Sledgehammer_Commands.default_params thy [])) | 
| 47060 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 wenzelm parents: 
47049diff
changeset | 304 | handle List.Empty => error "No ATP available." | 
| 33016 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 305 | in | 
| 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 306 | (case AList.lookup (op =) args proverK of | 
| 50352 | 307 | SOME name => name | 
| 308 | | NONE => default_prover_name ()) | |
| 309 | end | |
| 310 | ||
| 311 | fun get_prover ctxt name params goal all_facts = | |
| 312 | let | |
| 54503 | 313 | val learn = Sledgehammer_MaSh.mash_learn_proof ctxt params (prop_of goal) all_facts | 
| 50352 | 314 | in | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55201diff
changeset | 315 | Sledgehammer_Prover_Minimize.get_minimizing_prover ctxt Sledgehammer_Prover.Normal learn name | 
| 33016 
b73b74fe23c3
proper exceptions instead of unhandled partiality
 boehmes parents: 
32991diff
changeset | 316 | end | 
| 32525 | 317 | |
| 46340 | 318 | type stature = ATP_Problem_Generate.stature | 
| 38752 
6628adcae4a7
consider "locality" when assigning weights to facts
 blanchet parents: 
38700diff
changeset | 319 | |
| 58473 | 320 | fun is_good_line s = | 
| 53081 
2a62d848a56a
improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
 blanchet parents: 
52555diff
changeset | 321 | (String.isSubstring " ms)" s orelse String.isSubstring " s)" s) | 
| 
2a62d848a56a
improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
 blanchet parents: 
52555diff
changeset | 322 | andalso not (String.isSubstring "(> " s) | 
| 
2a62d848a56a
improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
 blanchet parents: 
52555diff
changeset | 323 | andalso not (String.isSubstring ", > " s) | 
| 
2a62d848a56a
improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
 blanchet parents: 
52555diff
changeset | 324 | andalso not (String.isSubstring "may fail" s) | 
| 
2a62d848a56a
improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
 blanchet parents: 
52555diff
changeset | 325 | |
| 51206 | 326 | (* Fragile hack *) | 
| 55285 | 327 | fun proof_method_from_msg args msg = | 
| 328 | (case AList.lookup (op =) args proof_methodK of | |
| 58473 | 329 | SOME name => | 
| 330 | if name = "smart" then | |
| 331 | if exists is_good_line (split_lines msg) then | |
| 332 | "none" | |
| 333 | else | |
| 334 | "fail" | |
| 335 | else | |
| 336 | name | |
| 41357 
ae76960d86a2
added "sledgehammer_tac" as possible reconstructor in Mirabelle
 blanchet parents: 
41338diff
changeset | 337 | | NONE => | 
| 58473 | 338 | if exists is_good_line (split_lines msg) then | 
| 51203 | 339 | "none" (* trust the preplayed proof *) | 
| 340 |     else if String.isSubstring "metis (" msg then
 | |
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 341 | msg |> Substring.full | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 342 |           |> Substring.position "metis ("
 | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 343 | |> snd |> Substring.position ")" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 344 | |> fst |> Substring.string | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 345 | |> suffix ")" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 346 | else if String.isSubstring "metis" msg then | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 347 | "metis" | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 348 | else | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 349 | "smt") | 
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 350 | |
| 32521 | 351 | local | 
| 352 | ||
| 32536 | 353 | datatype sh_result = | 
| 46340 | 354 | SH_OK of int * int * (string * stature) list | | 
| 32536 | 355 | SH_FAIL of int * int | | 
| 356 | SH_ERROR | |
| 357 | ||
| 50352 | 358 | fun run_sh prover_name fact_filter type_enc strict max_facts slice | 
| 50334 | 359 | lam_trans uncurried_aliases e_selection_heuristic term_order force_sos | 
| 58473 | 360 | hard_timeout timeout preplay_timeout isar_proofsLST smt_proofsLST | 
| 361 | minimizeLST max_new_mono_instancesLST max_mono_itersLST dir pos st = | |
| 32521 | 362 | let | 
| 55205 | 363 | val thy = Proof.theory_of st | 
| 38998 | 364 |     val {context = ctxt, facts = chained_ths, goal} = Proof.goal st
 | 
| 365 | val i = 1 | |
| 44090 | 366 | fun set_file_name (SOME dir) = | 
| 55212 | 367 | Config.put Sledgehammer_Prover_ATP.atp_dest_dir dir | 
| 368 | #> Config.put Sledgehammer_Prover_ATP.atp_problem_prefix | |
| 44424 | 369 |           ("prob_" ^ str0 (Position.line_of pos) ^ "__")
 | 
| 58061 | 370 | #> Config.put SMT_Config.debug_files | 
| 56811 | 371 | (dir ^ "/" ^ Name.desymbolize (SOME false) (ATP_Util.timestamp ()) ^ "_" | 
| 41338 | 372 | ^ serial_string ()) | 
| 44090 | 373 | | set_file_name NONE = I | 
| 39321 | 374 | val st' = | 
| 47030 | 375 | st | 
| 376 | |> Proof.map_context | |
| 377 | (set_file_name dir | |
| 47032 | 378 | #> (Option.map (Config.put ATP_Systems.e_selection_heuristic) | 
| 379 | e_selection_heuristic |> the_default I) | |
| 47049 | 380 | #> (Option.map (Config.put ATP_Systems.term_order) | 
| 381 | term_order |> the_default I) | |
| 47030 | 382 | #> (Option.map (Config.put ATP_Systems.force_sos) | 
| 383 | force_sos |> the_default I)) | |
| 57782 | 384 |     val params as {max_facts, minimize, preplay_timeout, ...} =
 | 
| 55205 | 385 | Sledgehammer_Commands.default_params thy | 
| 58473 | 386 |          ([(* ("verbose", "true"), *)
 | 
| 50334 | 387 |            ("fact_filter", fact_filter),
 | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43569diff
changeset | 388 |            ("type_enc", type_enc),
 | 
| 46386 | 389 |            ("strict", strict),
 | 
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 390 |            ("lam_trans", lam_trans |> the_default lam_trans_default),
 | 
| 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 391 |            ("uncurried_aliases", uncurried_aliases |> the_default uncurried_aliases_default),
 | 
| 48293 | 392 |            ("max_facts", max_facts),
 | 
| 45706 | 393 |            ("slice", slice),
 | 
| 44448 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 394 |            ("timeout", string_of_int timeout),
 | 
| 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 395 |            ("preplay_timeout", preplay_timeout)]
 | 
| 56411 | 396 | |> isar_proofsLST | 
| 58473 | 397 | |> smt_proofsLST | 
| 57782 | 398 | |> minimizeLST (*don't confuse the two minimization flags*) | 
| 46826 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 399 | |> max_new_mono_instancesLST | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 400 | |> max_mono_itersLST) | 
| 55205 | 401 | val default_max_facts = | 
| 402 | Sledgehammer_Prover_Minimize.default_max_facts_of_prover ctxt prover_name | |
| 52196 
2281f33e8da6
redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
 blanchet parents: 
52125diff
changeset | 403 | val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal i ctxt | 
| 32574 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 404 | 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 | 405 | (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 | 406 | NONE => I | 
| 
719426c9e1eb
added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
 boehmes parents: 
32571diff
changeset | 407 | | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs)) | 
| 42953 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 blanchet parents: 
42952diff
changeset | 408 | fun failed failure = | 
| 51009 
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
 blanchet parents: 
51008diff
changeset | 409 |       ({outcome = SOME failure, used_facts = [], used_from = [],
 | 
| 57751 
f453bbc289fa
modernized Mirabelle (a bit) and made it compile
 blanchet parents: 
57154diff
changeset | 410 | preferred_methss = (Sledgehammer_Proof_Methods.Auto_Method, []), run_time = Time.zeroTime, | 
| 
f453bbc289fa
modernized Mirabelle (a bit) and made it compile
 blanchet parents: 
57154diff
changeset | 411 | message = K ""}, ~1) | 
| 
f453bbc289fa
modernized Mirabelle (a bit) and made it compile
 blanchet parents: 
57154diff
changeset | 412 |     val ({outcome, used_facts, preferred_methss, run_time, message, ...}
 | 
| 55201 | 413 | : Sledgehammer_Prover.prover_result, | 
| 51009 
e8ff34a1fa9a
thread through fact triple component from which used facts come, for accurate index output
 blanchet parents: 
51008diff
changeset | 414 | time_isa) = time_limit (Mirabelle.cpu_time (fn () => | 
| 41275 | 415 | let | 
| 55212 | 416 | val ho_atp = Sledgehammer_Prover_ATP.is_ho_atp ctxt prover_name | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58921diff
changeset | 417 | val keywords = Thy_Header.get_keywords' ctxt | 
| 48299 | 418 | val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt | 
| 41275 | 419 | val facts = | 
| 48288 
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
 blanchet parents: 
48250diff
changeset | 420 | Sledgehammer_Fact.nearly_all_facts ctxt ho_atp | 
| 58921 
ffdafc99f67b
prefer explicit Keyword.keywords (cf. 82a71046dce8);
 wenzelm parents: 
58903diff
changeset | 421 | Sledgehammer_Fact.no_fact_override keywords css_table chained_ths | 
| 48299 | 422 | hyp_ts concl_t | 
| 51010 | 423 | val factss = | 
| 51007 
4f694d52bf62
thread fact triple (MeSh, MePo, MaSh) to allow different filters in different slices
 blanchet parents: 
51005diff
changeset | 424 | facts | 
| 48381 | 425 | |> Sledgehammer_MaSh.relevant_facts ctxt params prover_name | 
| 48293 | 426 | (the_default default_max_facts max_facts) | 
| 48292 | 427 | Sledgehammer_Fact.no_fact_override hyp_ts concl_t | 
| 51010 | 428 | |> tap (fn factss => | 
| 50868 | 429 | "Line " ^ str0 (Position.line_of pos) ^ ": " ^ | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55201diff
changeset | 430 | Sledgehammer.string_of_factss factss | 
| 58843 | 431 | |> writeln) | 
| 50352 | 432 | val prover = get_prover ctxt prover_name params goal facts | 
| 41275 | 433 | val problem = | 
| 54141 
f57f8e7a879f
generate a comment storing the goal nickname in "learn_prover"
 blanchet parents: 
54130diff
changeset | 434 |           {comment = "", state = st', goal = goal, subgoal = i,
 | 
| 51010 | 435 | subgoal_count = Sledgehammer_Util.subgoal_count st, factss = factss} | 
| 57751 
f453bbc289fa
modernized Mirabelle (a bit) and made it compile
 blanchet parents: 
57154diff
changeset | 436 | in prover params problem end)) () | 
| 42953 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 blanchet parents: 
42952diff
changeset | 437 | handle TimeLimit.TimeOut => failed ATP_Proof.TimedOut | 
| 
26111aafab12
detect inappropriate problems and crashes better in Waldmeister
 blanchet parents: 
42952diff
changeset | 438 | | Fail "inappropriate" => failed ATP_Proof.Inappropriate | 
| 45371 | 439 | val time_prover = run_time |> Time.toMilliseconds | 
| 57782 | 440 | val msg = message (fn () => Sledgehammer.play_one_line_proof minimize preplay_timeout used_facts | 
| 441 | st' i preferred_methss) | |
| 32521 | 442 | in | 
| 55211 | 443 | (case outcome of | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 444 | NONE => (msg, SH_OK (time_isa, time_prover, used_facts)) | 
| 55211 | 445 | | SOME _ => (msg, SH_FAIL (time_isa, time_prover))) | 
| 32521 | 446 | end | 
| 37994 
b04307085a09
make TPTP generator accept full first-order formulas
 blanchet parents: 
37631diff
changeset | 447 |   handle ERROR msg => ("error: " ^ msg, SH_ERROR)
 | 
| 32521 | 448 | |
| 32498 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32496diff
changeset | 449 | in | 
| 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32496diff
changeset | 450 | |
| 55285 | 451 | fun run_sledgehammer trivial args proof_method named_thms id | 
| 44090 | 452 |       ({pre=st, log, pos, ...}: Mirabelle.run_args) =
 | 
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: diff
changeset | 453 | let | 
| 55205 | 454 | val thy = Proof.theory_of st | 
| 39340 | 455 | val triv_str = if trivial then "[T] " else "" | 
| 32536 | 456 | val _ = change_data id inc_sh_calls | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 457 | val _ = if trivial then () else change_data id inc_sh_nontriv_calls | 
| 55205 | 458 | val prover_name = get_prover_name thy args | 
| 50334 | 459 | val fact_filter = AList.lookup (op =) args fact_filterK |> the_default fact_filter_default | 
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 460 | val type_enc = AList.lookup (op =) args type_encK |> the_default type_enc_default | 
| 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 461 | val strict = AList.lookup (op =) args strictK |> the_default strict_default | 
| 48293 | 462 | val max_facts = | 
| 55211 | 463 | (case AList.lookup (op =) args max_factsK of | 
| 48293 | 464 | SOME max => max | 
| 55211 | 465 | | NONE => | 
| 466 | (case AList.lookup (op =) args max_relevantK of | |
| 467 | SOME max => max | |
| 468 | | NONE => max_facts_default)) | |
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 469 | val slice = AList.lookup (op =) args sliceK |> the_default slice_default | 
| 45514 | 470 | val lam_trans = AList.lookup (op =) args lam_transK | 
| 46415 | 471 | val uncurried_aliases = AList.lookup (op =) args uncurried_aliasesK | 
| 47032 | 472 | val e_selection_heuristic = AList.lookup (op =) args e_selection_heuristicK | 
| 47049 | 473 | val term_order = AList.lookup (op =) args term_orderK | 
| 44099 | 474 | val force_sos = AList.lookup (op =) args force_sosK | 
| 42725 
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
 blanchet parents: 
42642diff
changeset | 475 | |> Option.map (curry (op <>) "false") | 
| 32525 | 476 | val dir = AList.lookup (op =) args keepK | 
| 32541 | 477 | val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30) | 
| 41268 | 478 | (* always use a hard timeout, but give some slack so that the automatic | 
| 479 | minimizer has a chance to do its magic *) | |
| 46825 
98300d5f9cc0
added sh_minimize and preplay_timeout options to Mirabelle's Sledgehammer action;
 sultana parents: 
46641diff
changeset | 480 | val preplay_timeout = AList.lookup (op =) args preplay_timeoutK | 
| 
98300d5f9cc0
added sh_minimize and preplay_timeout options to Mirabelle's Sledgehammer action;
 sultana parents: 
46641diff
changeset | 481 | |> the_default preplay_timeout_default | 
| 56411 | 482 | val isar_proofsLST = available_parameter args isar_proofsK "isar_proofs" | 
| 58473 | 483 | val smt_proofsLST = available_parameter args smt_proofsK "smt_proofs" | 
| 57782 | 484 | val minimizeLST = available_parameter args minimizeK "minimize" | 
| 46826 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 485 | val max_new_mono_instancesLST = | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 486 | available_parameter args max_new_mono_instancesK max_new_mono_instancesK | 
| 
4c80c4034f1d
added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
 sultana parents: 
46825diff
changeset | 487 | val max_mono_itersLST = available_parameter args max_mono_itersK max_mono_itersK | 
| 48702 
e1752ccccc34
increase defensive timeout that should typically not kick in thanks to solid time limits in Sledgehammer (otherwise, Mirabelle-based evaluations might be distorted)
 boehmes parents: 
48558diff
changeset | 488 | val hard_timeout = SOME (4 * timeout) | 
| 41155 
85da8cbb4966
added support for "type_sys" option to Mirabelle
 blanchet parents: 
41154diff
changeset | 489 | val (msg, result) = | 
| 50352 | 490 | run_sh prover_name fact_filter type_enc strict max_facts slice lam_trans | 
| 491 | uncurried_aliases e_selection_heuristic term_order force_sos | |
| 58473 | 492 | hard_timeout timeout preplay_timeout isar_proofsLST smt_proofsLST | 
| 493 | minimizeLST max_new_mono_instancesLST max_mono_itersLST dir pos st | |
| 32525 | 494 | in | 
| 55211 | 495 | (case result of | 
| 40062 | 496 | SH_OK (time_isa, time_prover, names) => | 
| 38700 | 497 | let | 
| 46340 | 498 | fun get_thms (name, stature) = | 
| 50267 
1da2e67242d6
moved thms_of_name to Sledgehammer_Util and removed copies, updated references
 smolkas parents: 
49916diff
changeset | 499 | try (Sledgehammer_Util.thms_of_name (Proof.context_of st)) | 
| 49916 
412346127bfa
fixed theorem lookup code in Isar proof reconstruction
 blanchet parents: 
49914diff
changeset | 500 | name | 
| 47154 
2c357e2b8436
more robustness in case a theorem cannot be retrieved (which typically happens with backtick facts)
 blanchet parents: 
47060diff
changeset | 501 | |> Option.map (pair (name, stature)) | 
| 32525 | 502 | in | 
| 32818 | 503 | change_data id inc_sh_success; | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 504 | if trivial then () else change_data id inc_sh_nontriv_success; | 
| 32818 | 505 | change_data id (inc_sh_lemmas (length names)); | 
| 506 | change_data id (inc_sh_max_lems (length names)); | |
| 507 | change_data id (inc_sh_time_isa time_isa); | |
| 40062 | 508 | change_data id (inc_sh_time_prover time_prover); | 
| 55285 | 509 | proof_method := proof_method_from_msg args msg; | 
| 38826 | 510 | named_thms := SOME (map_filter get_thms names); | 
| 39340 | 511 |           log (sh_tag id ^ triv_str ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
 | 
| 40062 | 512 | string_of_int time_prover ^ ") [" ^ prover_name ^ "]:\n" ^ msg) | 
| 32525 | 513 | end | 
| 40062 | 514 | | SH_FAIL (time_isa, time_prover) => | 
| 32536 | 515 | let | 
| 516 | val _ = change_data id (inc_sh_time_isa time_isa) | |
| 40062 | 517 | val _ = change_data id (inc_sh_time_prover_fail time_prover) | 
| 39340 | 518 | in log (sh_tag id ^ triv_str ^ "failed: " ^ msg) end | 
| 55211 | 519 | | SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)) | 
| 32525 | 520 | end | 
| 521 | ||
| 522 | end | |
| 523 | ||
| 44542 | 524 | fun override_params prover type_enc timeout = | 
| 525 |   [("provers", prover),
 | |
| 48293 | 526 |    ("max_facts", "0"),
 | 
| 44542 | 527 |    ("type_enc", type_enc),
 | 
| 46386 | 528 |    ("strict", "true"),
 | 
| 45706 | 529 |    ("slice", "false"),
 | 
| 44461 | 530 |    ("timeout", timeout |> Time.toSeconds |> string_of_int)]
 | 
| 44430 | 531 | |
| 57753 | 532 | fun run_proof_method trivial full name meth named_thms id | 
| 32567 
de411627a985
explicitly export type abbreviations (as usual in SML97);
 wenzelm parents: 
32564diff
changeset | 533 |     ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
 | 
| 32525 | 534 | let | 
| 55285 | 535 | fun do_method named_thms ctxt = | 
| 44462 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 blanchet parents: 
44461diff
changeset | 536 | let | 
| 59083 | 537 | val ref_of_str = (* FIXME proper wrapper for parser combinators *) | 
| 538 | suffix ";" #> Token.explode (Thy_Header.get_keywords' ctxt) Position.none | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58921diff
changeset | 539 | #> Parse.xthm #> fst | 
| 44462 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 blanchet parents: 
44461diff
changeset | 540 | val thms = named_thms |> maps snd | 
| 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 blanchet parents: 
44461diff
changeset | 541 | val facts = named_thms |> map (ref_of_str o fst o fst) | 
| 48292 | 542 |         val fact_override = {add = facts, del = [], only = true}
 | 
| 44566 
bf8331161ad9
split timeout among ATPs in and add Metis to the mix as backup
 blanchet parents: 
44542diff
changeset | 543 | fun my_timeout time_slice = | 
| 
bf8331161ad9
split timeout among ATPs in and add Metis to the mix as backup
 blanchet parents: 
44542diff
changeset | 544 | timeout |> Time.toReal |> curry Real.* time_slice |> Time.fromReal | 
| 
bf8331161ad9
split timeout among ATPs in and add Metis to the mix as backup
 blanchet parents: 
44542diff
changeset | 545 | fun sledge_tac time_slice prover type_enc = | 
| 44542 | 546 | Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt | 
| 57814 | 547 | (override_params prover type_enc (my_timeout time_slice)) fact_override [] | 
| 44462 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 blanchet parents: 
44461diff
changeset | 548 | in | 
| 55285 | 549 | if !meth = "sledgehammer_tac" then | 
| 57154 | 550 | sledge_tac 0.2 ATP_Proof.vampireN "mono_native" | 
| 551 | ORELSE' sledge_tac 0.2 ATP_Proof.eN "poly_guards??" | |
| 552 | ORELSE' sledge_tac 0.2 ATP_Proof.spassN "mono_native" | |
| 553 | ORELSE' sledge_tac 0.2 ATP_Proof.z3_tptpN "poly_tags??" | |
| 58061 | 554 | ORELSE' SMT_Solver.smt_tac ctxt thms | 
| 55285 | 555 | else if !meth = "smt" then | 
| 58061 | 556 | SMT_Solver.smt_tac ctxt thms | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 557 | else if full then | 
| 46320 | 558 | Metis_Tactic.metis_tac [ATP_Proof_Reconstruct.full_typesN] | 
| 54500 | 559 | ATP_Proof_Reconstruct.default_metis_lam_trans ctxt thms | 
| 55285 | 560 |         else if String.isPrefix "metis (" (!meth) then
 | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 561 | let | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 562 | val (type_encs, lam_trans) = | 
| 55285 | 563 | !meth | 
| 59083 | 564 | |> Token.explode (Thy_Header.get_keywords' ctxt) Position.start | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 565 | |> filter Token.is_proper |> tl | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 566 | |> Metis_Tactic.parse_metis_options |> fst | 
| 46320 | 567 | |>> the_default [ATP_Proof_Reconstruct.partial_typesN] | 
| 54500 | 568 | ||> the_default ATP_Proof_Reconstruct.default_metis_lam_trans | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 569 | in Metis_Tactic.metis_tac type_encs lam_trans ctxt thms end | 
| 55285 | 570 | else if !meth = "metis" then | 
| 55211 | 571 | Metis_Tactic.metis_tac [] ATP_Proof_Reconstruct.default_metis_lam_trans ctxt thms | 
| 58473 | 572 | else if !meth = "none" then | 
| 573 | K all_tac | |
| 574 | else if !meth = "fail" then | |
| 575 | K no_tac | |
| 44462 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 blanchet parents: 
44461diff
changeset | 576 | else | 
| 58473 | 577 |           (warning ("Unknown method " ^ quote (!meth)); K no_tac)
 | 
| 44462 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 blanchet parents: 
44461diff
changeset | 578 | end | 
| 55285 | 579 | fun apply_method named_thms = | 
| 580 | Mirabelle.can_apply timeout (do_method named_thms) st | |
| 32521 | 581 | |
| 582 |     fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
 | |
| 57753 | 583 | | with_time (true, t) = (change_data id inc_proof_method_success; | 
| 40667 
b8579f24ce67
make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
 blanchet parents: 
40627diff
changeset | 584 | if trivial then () | 
| 57753 | 585 | else change_data id inc_proof_method_nontriv_success; | 
| 586 | change_data id (inc_proof_method_lemmas (length named_thms)); | |
| 587 | change_data id (inc_proof_method_time t); | |
| 588 | change_data id (inc_proof_method_posns (pos, trivial)); | |
| 589 | if name = "proof" then change_data id inc_proof_method_proofs else (); | |
| 32521 | 590 |           "succeeded (" ^ string_of_int t ^ ")")
 | 
| 55285 | 591 | fun timed_method named_thms = | 
| 592 | (with_time (Mirabelle.cpu_time apply_method named_thms), true) | |
| 57753 | 593 |       handle TimeLimit.TimeOut => (change_data id inc_proof_method_timeout; ("timeout", false))
 | 
| 34052 | 594 |            | ERROR msg => ("error: " ^ msg, false)
 | 
| 32521 | 595 | |
| 32525 | 596 | val _ = log separator | 
| 57753 | 597 | val _ = change_data id inc_proof_method_calls | 
| 598 | val _ = if trivial then () else change_data id inc_proof_method_nontriv_calls | |
| 32521 | 599 | in | 
| 44462 
d9a657c44380
more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
 blanchet parents: 
44461diff
changeset | 600 | named_thms | 
| 55285 | 601 | |> timed_method | 
| 602 | |>> log o prefix (proof_method_tag meth id) | |
| 34052 | 603 | |> snd | 
| 32521 | 604 | end | 
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: diff
changeset | 605 | |
| 41276 
285aea0c153c
two layers of timeouts seem to be less reliable than a single layer
 blanchet parents: 
41275diff
changeset | 606 | val try_timeout = seconds 5.0 | 
| 39337 
ffa577c0bbfa
keep track of trivial vs. nontrivial calls using "try" for 30 seconds
 blanchet parents: 
39321diff
changeset | 607 | |
| 44431 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 blanchet parents: 
44430diff
changeset | 608 | (* crude hack *) | 
| 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 blanchet parents: 
44430diff
changeset | 609 | val num_sledgehammer_calls = Unsynchronized.ref 0 | 
| 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 blanchet parents: 
44430diff
changeset | 610 | |
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33316diff
changeset | 611 | 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: 
34052diff
changeset | 612 | 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: 
34052diff
changeset | 613 | 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: 
34052diff
changeset | 614 | then () else | 
| 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 wenzelm parents: 
34052diff
changeset | 615 | let | 
| 44431 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 blanchet parents: 
44430diff
changeset | 616 | val max_calls = | 
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 617 | AList.lookup (op =) args max_callsK |> the_default max_calls_default | 
| 44431 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 blanchet parents: 
44430diff
changeset | 618 | |> Int.fromString |> the | 
| 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 blanchet parents: 
44430diff
changeset | 619 | val _ = num_sledgehammer_calls := !num_sledgehammer_calls + 1; | 
| 44448 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 620 | in | 
| 44431 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 blanchet parents: 
44430diff
changeset | 621 | if !num_sledgehammer_calls > max_calls then () | 
| 
3e0ce0ae1022
added "max_calls" option to get a fixed number of Sledgehammer calls per theory
 blanchet parents: 
44430diff
changeset | 622 | else | 
| 44448 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 623 | let | 
| 55285 | 624 | val meth = Unsynchronized.ref "" | 
| 44448 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 625 | val named_thms = | 
| 46340 | 626 | Unsynchronized.ref (NONE : ((string * stature) * thm list) list option) | 
| 44448 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 627 | val trivial = | 
| 47481 
b5873d4ff1e2
gathered mirabelle_sledgehammer's hardcoded defaults
 sultana parents: 
47480diff
changeset | 628 | if AList.lookup (op =) args check_trivialK |> the_default trivial_default | 
| 51951 
fab4ab92e812
more standard Isabelle/ML operations -- avoid inaccurate Bool.fromString;
 wenzelm parents: 
51210diff
changeset | 629 | |> Markup.parse_bool then | 
| 47201 | 630 | Try0.try0 (SOME try_timeout) ([], [], [], []) pre | 
| 631 | handle TimeLimit.TimeOut => false | |
| 632 | else false | |
| 57753 | 633 | fun apply_method () = | 
| 634 | (Mirabelle.catch_result (proof_method_tag meth) false | |
| 635 | (run_proof_method trivial false name meth (these (!named_thms))) id st; ()) | |
| 44448 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 636 | in | 
| 55285 | 637 | Mirabelle.catch sh_tag (run_sledgehammer trivial args meth named_thms) id st; | 
| 57753 | 638 | if is_some (!named_thms) then apply_method () else () | 
| 44448 
04bd6a9307c6
don't perform a triviality check if the goal is skipped anyway
 blanchet parents: 
44447diff
changeset | 639 | end | 
| 35592 
768d17f54125
use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
 wenzelm parents: 
34052diff
changeset | 640 | end | 
| 32818 | 641 | end | 
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: diff
changeset | 642 | |
| 32511 | 643 | fun invoke args = | 
| 43569 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 blanchet parents: 
43351diff
changeset | 644 | Mirabelle.register (init, sledgehammer_action args, done) | 
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: diff
changeset | 645 | |
| 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: diff
changeset | 646 | end |