author | blanchet |
Wed, 18 Jul 2012 08:44:04 +0200 | |
changeset 48307 | 7c78f14d20cf |
parent 48303 | f1d135d0ea69 |
child 48311 | 3c4e10606567 |
permissions | -rw-r--r-- |
48285 | 1 |
(* Title: HOL/TPTP/mash_eval.ML |
48234 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
Copyright 2012 |
|
4 |
||
48285 | 5 |
Evaluate proof suggestions from MaSh (Machine-learning for Sledgehammer). |
48234 | 6 |
*) |
7 |
||
48285 | 8 |
signature MASH_EVAL = |
48234 | 9 |
sig |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
10 |
type params = Sledgehammer_Provers.params |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
11 |
|
48285 | 12 |
val evaluate_mash_suggestions : |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
13 |
Proof.context -> params -> theory -> string -> unit |
48234 | 14 |
end; |
15 |
||
48285 | 16 |
structure MaSh_Eval : MASH_EVAL = |
48234 | 17 |
struct |
48235 | 18 |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
19 |
open Sledgehammer_Filter_MaSh |
48240 | 20 |
|
48300 | 21 |
val isarN = "Isa" |
48285 | 22 |
val iterN = "Iter" |
48241 | 23 |
val mashN = "MaSh" |
48298 | 24 |
val meshN = "Mesh" |
48241 | 25 |
|
48293 | 26 |
val max_facts_slack = 2 |
48235 | 27 |
|
48285 | 28 |
fun evaluate_mash_suggestions ctxt params thy file_name = |
48235 | 29 |
let |
48293 | 30 |
val {provers, max_facts, slice, type_enc, lam_trans, timeout, ...} = |
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48236
diff
changeset
|
31 |
Sledgehammer_Isar.default_params ctxt [] |
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48236
diff
changeset
|
32 |
val prover_name = hd provers |
48235 | 33 |
val path = file_name |> Path.explode |
34 |
val lines = path |> File.read_lines |
|
48299 | 35 |
val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt |
36 |
val facts = all_non_tautological_facts_of thy css_table |
|
48235 | 37 |
val all_names = facts |> map (Thm.get_name_hint o snd) |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
38 |
val iter_ok = Unsynchronized.ref 0 |
48241 | 39 |
val mash_ok = Unsynchronized.ref 0 |
48298 | 40 |
val mesh_ok = Unsynchronized.ref 0 |
48300 | 41 |
val isar_ok = Unsynchronized.ref 0 |
48235 | 42 |
fun find_sugg facts sugg = |
43 |
find_first (fn (_, th) => Thm.get_name_hint th = sugg) facts |
|
48236 | 44 |
fun sugg_facts hyp_ts concl_t facts = |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48300
diff
changeset
|
45 |
map_filter (find_sugg facts) |
48293 | 46 |
#> take (max_facts_slack * the max_facts) |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
48245
diff
changeset
|
47 |
#> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts concl_t |
48298 | 48 |
fun mesh_facts fsp = |
48235 | 49 |
let |
48289 | 50 |
val (fs1, fs2) = |
51 |
fsp |> pairself (map (apfst (apfst (fn name => name ())))) |
|
48235 | 52 |
val fact_eq = (op =) o pairself fst |
53 |
fun score_in f fs = |
|
54 |
case find_index (curry fact_eq f) fs of |
|
55 |
~1 => length fs |
|
56 |
| j => j |
|
57 |
fun score_of f = score_in f fs1 + score_in f fs2 |
|
58 |
in |
|
59 |
union fact_eq fs1 fs2 |
|
48236 | 60 |
|> map (`score_of) |> sort (int_ord o pairself fst) |> map snd |
48293 | 61 |
|> take (the max_facts) |
48289 | 62 |
|> map (apfst (apfst K)) |
48235 | 63 |
end |
48289 | 64 |
fun with_index facts s = (find_index (curry (op =) s) facts + 1, s) |
48235 | 65 |
fun index_string (j, s) = s ^ "@" ^ string_of_int j |
48240 | 66 |
fun str_of_res label facts {outcome, run_time, used_facts, ...} = |
48289 | 67 |
let val facts = facts |> map (fn ((name, _), _) => name ()) in |
68 |
" " ^ label ^ ": " ^ |
|
69 |
(if is_none outcome then |
|
70 |
"Success (" ^ ATP_Util.string_from_time run_time ^ "): " ^ |
|
71 |
(used_facts |> map (with_index facts o fst) |
|
72 |
|> sort (int_ord o pairself fst) |
|
73 |
|> map index_string |
|
74 |
|> space_implode " ") ^ |
|
48293 | 75 |
(if length facts < the max_facts then |
48289 | 76 |
" (of " ^ string_of_int (length facts) ^ ")" |
77 |
else |
|
78 |
"") |
|
79 |
else |
|
80 |
"Failure: " ^ |
|
48293 | 81 |
(facts |> take (the max_facts) |> tag_list 1 |
48289 | 82 |
|> map index_string |
83 |
|> space_implode " ")) |
|
84 |
end |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
85 |
fun prove ok heading facts goal = |
48239
0016290f904c
generate Meng--Paulson facts for evaluation purposes
blanchet
parents:
48236
diff
changeset
|
86 |
let |
48293 | 87 |
val facts = facts |> take (the max_facts) |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
88 |
val res as {outcome, ...} = run_prover ctxt params facts goal |
48241 | 89 |
val _ = if is_none outcome then ok := !ok + 1 else () |
90 |
in str_of_res heading facts res end |
|
91 |
fun solve_goal j name suggs = |
|
48235 | 92 |
let |
93 |
val th = |
|
94 |
case find_first (fn (_, th) => Thm.get_name_hint th = name) facts of |
|
95 |
SOME (_, th) => th |
|
96 |
| NONE => error ("No fact called \"" ^ name ^ "\"") |
|
48300 | 97 |
val isar_deps = isabelle_dependencies_of all_names th |
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
48245
diff
changeset
|
98 |
val goal = goal_of_thm thy th |
48245 | 99 |
val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1 |
48235 | 100 |
val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS) |
48300 | 101 |
val isar_facts = sugg_facts hyp_ts concl_t facts isar_deps |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
102 |
val iter_facts = |
48292 | 103 |
Sledgehammer_Filter_Iter.iterative_relevant_facts ctxt params |
48293 | 104 |
prover_name (max_facts_slack * the max_facts) NONE hyp_ts concl_t |
105 |
facts |
|
48236 | 106 |
val mash_facts = sugg_facts hyp_ts concl_t facts suggs |
48298 | 107 |
val mesh_facts = mesh_facts (iter_facts, mash_facts) |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
108 |
val iter_s = prove iter_ok iterN iter_facts goal |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
109 |
val mash_s = prove mash_ok mashN mash_facts goal |
48298 | 110 |
val mesh_s = prove mesh_ok meshN mesh_facts goal |
48300 | 111 |
val isar_s = prove isar_ok isarN isar_facts goal |
48235 | 112 |
in |
48300 | 113 |
["Goal " ^ string_of_int j ^ ": " ^ name, iter_s, mash_s, mesh_s, |
114 |
isar_s] |
|
48241 | 115 |
|> cat_lines |> tracing |
48235 | 116 |
end |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48300
diff
changeset
|
117 |
val explode_suggs = |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48300
diff
changeset
|
118 |
space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta |
48241 | 119 |
fun do_line (j, line) = |
48235 | 120 |
case space_explode ":" line of |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48300
diff
changeset
|
121 |
[goal_name, suggs] => |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48300
diff
changeset
|
122 |
solve_goal j (unescape_meta goal_name) (explode_suggs suggs) |
48235 | 123 |
| _ => () |
48241 | 124 |
fun total_of heading ok n = |
125 |
" " ^ heading ^ ": " ^ string_of_int (!ok) ^ " (" ^ |
|
126 |
Real.fmt (StringCvt.FIX (SOME 1)) |
|
127 |
(100.0 * Real.fromInt (!ok) / Real.fromInt n) ^ "%)" |
|
48250
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
48245
diff
changeset
|
128 |
val inst_inducts = Config.get ctxt Sledgehammer_Fact.instantiate_inducts |
48245 | 129 |
val options = |
48293 | 130 |
[prover_name, string_of_int (the max_facts) ^ " facts", |
48241 | 131 |
"slice" |> not slice ? prefix "dont_", the_default "smart" type_enc, |
132 |
the_default "smart" lam_trans, ATP_Util.string_from_time timeout, |
|
133 |
"instantiate_inducts" |> not inst_inducts ? prefix "dont_"] |
|
134 |
val n = length lines |
|
135 |
in |
|
136 |
tracing " * * *"; |
|
48245 | 137 |
tracing ("Options: " ^ commas options); |
48241 | 138 |
List.app do_line (tag_list 1 lines); |
139 |
["Successes (of " ^ string_of_int n ^ " goals)", |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48250
diff
changeset
|
140 |
total_of iterN iter_ok n, |
48241 | 141 |
total_of mashN mash_ok n, |
48298 | 142 |
total_of meshN mesh_ok n, |
48300 | 143 |
total_of isarN isar_ok n] |
48241 | 144 |
|> cat_lines |> tracing |
145 |
end |
|
48235 | 146 |
|
48234 | 147 |
end; |