author | blanchet |
Fri, 20 Jul 2012 22:19:46 +0200 | |
changeset 48384 | 83dc102041e6 |
parent 48383 | df75b2d7e26a |
child 48385 | 2779dea0b1e0 |
permissions | -rw-r--r-- |
48380 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_mash.ML |
48248 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
||
4 |
Sledgehammer's machine-learning-based relevance filter (MaSh). |
|
5 |
*) |
|
6 |
||
48381 | 7 |
signature SLEDGEHAMMER_MASH = |
48248 | 8 |
sig |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
9 |
type status = ATP_Problem_Generate.status |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
10 |
type stature = ATP_Problem_Generate.stature |
48296
e7f01b7e244e
gracefully handle the case of empty theories when going up the accessibility chain
blanchet
parents:
48293
diff
changeset
|
11 |
type fact = Sledgehammer_Fact.fact |
e7f01b7e244e
gracefully handle the case of empty theories when going up the accessibility chain
blanchet
parents:
48293
diff
changeset
|
12 |
type fact_override = Sledgehammer_Fact.fact_override |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
13 |
type params = Sledgehammer_Provers.params |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
14 |
type relevance_fudge = Sledgehammer_Provers.relevance_fudge |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
15 |
type prover_result = Sledgehammer_Provers.prover_result |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
16 |
|
48308 | 17 |
val trace : bool Config.T |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
18 |
val MaShN : string |
48379
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
19 |
val mepoN : string |
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
20 |
val mashN : string |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
21 |
val meshN : string |
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
22 |
val fact_filters : string list |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
23 |
val escape_meta : string -> string |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
24 |
val escape_metas : string list -> string |
48308 | 25 |
val unescape_meta : string -> string |
26 |
val unescape_metas : string -> string list |
|
48311 | 27 |
val extract_query : string -> string * string list |
48378 | 28 |
val nickname_of : thm -> string |
48321 | 29 |
val suggested_facts : string list -> ('a * thm) list -> ('a * thm) list |
30 |
val mesh_facts : |
|
31 |
int -> (('a * thm) list * ('a * thm) list) list -> ('a * thm) list |
|
48324 | 32 |
val is_likely_tautology_or_too_meta : thm -> bool |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
33 |
val theory_ord : theory * theory -> order |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
34 |
val thm_ord : thm * thm -> order |
48318 | 35 |
val features_of : |
36 |
Proof.context -> string -> theory -> status -> term list -> string list |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
37 |
val isabelle_dependencies_of : unit Symtab.table -> thm -> string list |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
38 |
val goal_of_thm : theory -> thm -> thm |
48321 | 39 |
val run_prover_for_mash : |
48318 | 40 |
Proof.context -> params -> string -> fact list -> thm -> prover_result |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
41 |
val mash_CLEAR : Proof.context -> unit |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
42 |
val mash_INIT : |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
43 |
Proof.context -> bool |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
44 |
-> (string * string list * string list * string list) list -> unit |
48308 | 45 |
val mash_ADD : |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
46 |
Proof.context -> bool |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
47 |
-> (string * string list * string list * string list) list -> unit |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
48 |
val mash_QUERY : |
48318 | 49 |
Proof.context -> bool -> int -> string list * string list -> string list |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
50 |
val mash_unlearn : Proof.context -> unit |
48318 | 51 |
val mash_could_suggest_facts : unit -> bool |
48321 | 52 |
val mash_can_suggest_facts : Proof.context -> bool |
48298 | 53 |
val mash_suggest_facts : |
48321 | 54 |
Proof.context -> params -> string -> int -> term list -> term |
55 |
-> ('a * thm) list -> ('a * thm) list * ('a * thm) list |
|
48383 | 56 |
val mash_learn_proof : |
48384
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
57 |
Proof.context -> params -> string -> term -> ('a * thm) list -> thm list |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
58 |
-> unit |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
59 |
val mash_learn_thy : |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
60 |
Proof.context -> params -> theory -> Time.time -> fact list -> string |
48383 | 61 |
val mash_learn : Proof.context -> params -> unit |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
62 |
val relevant_facts : |
48292 | 63 |
Proof.context -> params -> string -> int -> fact_override -> term list |
48296
e7f01b7e244e
gracefully handle the case of empty theories when going up the accessibility chain
blanchet
parents:
48293
diff
changeset
|
64 |
-> term -> fact list -> fact list |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
65 |
val kill_learners : unit -> unit |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
66 |
val running_learners : unit -> unit |
48248 | 67 |
end; |
68 |
||
48381 | 69 |
structure Sledgehammer_MaSh : SLEDGEHAMMER_MASH = |
48248 | 70 |
struct |
48249 | 71 |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
72 |
open ATP_Util |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
73 |
open ATP_Problem_Generate |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
74 |
open Sledgehammer_Util |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
75 |
open Sledgehammer_Fact |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
76 |
open Sledgehammer_Provers |
48318 | 77 |
open Sledgehammer_Minimize |
48381 | 78 |
open Sledgehammer_MePo |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
79 |
|
48308 | 80 |
val trace = |
48380 | 81 |
Attrib.setup_config_bool @{binding sledgehammer_mash_trace} (K false) |
48308 | 82 |
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else () |
83 |
||
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
84 |
val MaShN = "MaSh" |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
85 |
|
48379
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
86 |
val mepoN = "mepo" |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
87 |
val mashN = "mash" |
48379
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
88 |
val meshN = "mesh" |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
89 |
|
48379
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
90 |
val fact_filters = [meshN, mepoN, mashN] |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
91 |
|
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
92 |
fun mash_home () = getenv "MASH_HOME" |
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
93 |
fun mash_state_dir () = |
48309 | 94 |
getenv "ISABELLE_HOME_USER" ^ "/mash" |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
95 |
|> tap (Isabelle_System.mkdir o Path.explode) |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
96 |
fun mash_state_path () = mash_state_dir () ^ "/state" |> Path.explode |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
97 |
|
48330 | 98 |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
99 |
(*** Isabelle helpers ***) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
100 |
|
48308 | 101 |
fun meta_char c = |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
102 |
if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
103 |
c = #")" orelse c = #"," then |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
104 |
String.str c |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
105 |
else |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
106 |
(* fixed width, in case more digits follow *) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
107 |
"\\" ^ stringN_of_int 3 (Char.ord c) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
108 |
|
48308 | 109 |
fun unmeta_chars accum [] = String.implode (rev accum) |
110 |
| unmeta_chars accum (#"\\" :: d1 :: d2 :: d3 :: cs) = |
|
111 |
(case Int.fromString (String.implode [d1, d2, d3]) of |
|
112 |
SOME n => unmeta_chars (Char.chr n :: accum) cs |
|
113 |
| NONE => "" (* error *)) |
|
114 |
| unmeta_chars _ (#"\\" :: _) = "" (* error *) |
|
115 |
| unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs |
|
116 |
||
117 |
val escape_meta = String.translate meta_char |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
118 |
val escape_metas = map escape_meta #> space_implode " " |
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48314
diff
changeset
|
119 |
val unescape_meta = String.explode #> unmeta_chars [] |
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48314
diff
changeset
|
120 |
val unescape_metas = |
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48314
diff
changeset
|
121 |
space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
122 |
|
48311 | 123 |
fun extract_query line = |
124 |
case space_explode ":" line of |
|
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48314
diff
changeset
|
125 |
[goal_name, suggs] => (unescape_meta goal_name, unescape_metas suggs) |
48312 | 126 |
| _ => ("", []) |
48311 | 127 |
|
48378 | 128 |
fun parent_of_local_thm th = |
129 |
let |
|
130 |
val thy = th |> Thm.theory_of_thm |
|
131 |
val facts = thy |> Global_Theory.facts_of |
|
132 |
val space = facts |> Facts.space_of |
|
133 |
fun id_of s = #id (Name_Space.the_entry space s) |
|
134 |
fun max_id (s', _) (s, id) = |
|
135 |
let val id' = id_of s' in if id > id' then (s, id) else (s', id') end |
|
136 |
in ("", ~1) |> Facts.fold_static max_id facts |> fst end |
|
137 |
||
138 |
val local_prefix = "local" ^ Long_Name.separator |
|
139 |
||
140 |
fun nickname_of th = |
|
141 |
let val hint = Thm.get_name_hint th in |
|
142 |
(* FIXME: There must be a better way to detect local facts. *) |
|
143 |
case try (unprefix local_prefix) hint of |
|
144 |
SOME suff => |
|
145 |
parent_of_local_thm th ^ Long_Name.separator ^ Long_Name.separator ^ suff |
|
146 |
| NONE => hint |
|
147 |
end |
|
148 |
||
48330 | 149 |
fun suggested_facts suggs facts = |
150 |
let |
|
48378 | 151 |
fun add_fact (fact as (_, th)) = Symtab.default (nickname_of th, fact) |
48330 | 152 |
val tab = Symtab.empty |> fold add_fact facts |
153 |
in map_filter (Symtab.lookup tab) suggs end |
|
48311 | 154 |
|
48329
5781c4620245
use better score function, based on previous evaluation (cf. Deduct 2011 slides)
blanchet
parents:
48328
diff
changeset
|
155 |
(* Ad hoc score function roughly based on Blanchette's Ringberg 2011 data. *) |
5781c4620245
use better score function, based on previous evaluation (cf. Deduct 2011 slides)
blanchet
parents:
48328
diff
changeset
|
156 |
fun score x = Math.pow (1.5, 15.5 - 0.05 * Real.fromInt x) + 15.0 |
48328
ca0b7d19dd62
attempt at meshing according to more meaningful factors
blanchet
parents:
48327
diff
changeset
|
157 |
|
ca0b7d19dd62
attempt at meshing according to more meaningful factors
blanchet
parents:
48327
diff
changeset
|
158 |
fun sum_sq_avg [] = 0 |
48329
5781c4620245
use better score function, based on previous evaluation (cf. Deduct 2011 slides)
blanchet
parents:
48328
diff
changeset
|
159 |
| sum_sq_avg xs = |
5781c4620245
use better score function, based on previous evaluation (cf. Deduct 2011 slides)
blanchet
parents:
48328
diff
changeset
|
160 |
Real.ceil (100000.0 * fold (curry (op +) o score) xs 0.0) div length xs |
48313
0faafdffa662
mesh facts by taking into consideration whether a fact is known to MeSh
blanchet
parents:
48312
diff
changeset
|
161 |
|
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
162 |
fun mesh_facts max_facts [(selected, unknown)] = |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
163 |
take max_facts selected @ take (max_facts - length selected) unknown |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
164 |
| mesh_facts max_facts mess = |
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
165 |
let |
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
166 |
val mess = mess |> map (apfst (`length)) |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
167 |
val fact_eq = Thm.eq_thm o pairself snd |
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
168 |
fun score_in fact ((sel_len, sels), unks) = |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
169 |
case find_index (curry fact_eq fact) sels of |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
170 |
~1 => (case find_index (curry fact_eq fact) unks of |
48329
5781c4620245
use better score function, based on previous evaluation (cf. Deduct 2011 slides)
blanchet
parents:
48328
diff
changeset
|
171 |
~1 => SOME sel_len |
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
172 |
| _ => NONE) |
48329
5781c4620245
use better score function, based on previous evaluation (cf. Deduct 2011 slides)
blanchet
parents:
48328
diff
changeset
|
173 |
| j => SOME j |
48328
ca0b7d19dd62
attempt at meshing according to more meaningful factors
blanchet
parents:
48327
diff
changeset
|
174 |
fun score_of fact = mess |> map_filter (score_in fact) |> sum_sq_avg |
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
175 |
val facts = fold (union fact_eq o take max_facts o snd o fst) mess [] |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
176 |
in |
48328
ca0b7d19dd62
attempt at meshing according to more meaningful factors
blanchet
parents:
48327
diff
changeset
|
177 |
facts |> map (`score_of) |> sort (int_ord o swap o pairself fst) |
ca0b7d19dd62
attempt at meshing according to more meaningful factors
blanchet
parents:
48327
diff
changeset
|
178 |
|> map snd |> take max_facts |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
179 |
end |
48312 | 180 |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
181 |
val thy_feature_prefix = "y_" |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
182 |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
183 |
val thy_feature_name_of = prefix thy_feature_prefix |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
184 |
val const_name_of = prefix const_prefix |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
185 |
val type_name_of = prefix type_const_prefix |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
186 |
val class_name_of = prefix class_prefix |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
187 |
|
48324 | 188 |
fun is_likely_tautology_or_too_meta th = |
189 |
let |
|
190 |
val is_boring_const = member (op =) atp_widely_irrelevant_consts |
|
191 |
fun is_boring_bool t = |
|
192 |
not (exists_Const (not o is_boring_const o fst) t) orelse |
|
193 |
exists_type (exists_subtype (curry (op =) @{typ prop})) t |
|
194 |
fun is_boring_prop (@{const Trueprop} $ t) = is_boring_bool t |
|
195 |
| is_boring_prop (@{const "==>"} $ t $ u) = |
|
196 |
is_boring_prop t andalso is_boring_prop u |
|
197 |
| is_boring_prop (Const (@{const_name all}, _) $ (Abs (_, _, t)) $ u) = |
|
198 |
is_boring_prop t andalso is_boring_prop u |
|
199 |
| is_boring_prop (Const (@{const_name "=="}, _) $ t $ u) = |
|
200 |
is_boring_bool t andalso is_boring_bool u |
|
201 |
| is_boring_prop _ = true |
|
202 |
in |
|
203 |
is_boring_prop (prop_of th) andalso not (Thm.eq_thm_prop (@{thm ext}, th)) |
|
204 |
end |
|
205 |
||
206 |
fun theory_ord p = |
|
207 |
if Theory.eq_thy p then |
|
208 |
EQUAL |
|
209 |
else if Theory.subthy p then |
|
210 |
LESS |
|
211 |
else if Theory.subthy (swap p) then |
|
212 |
GREATER |
|
213 |
else case int_ord (pairself (length o Theory.ancestors_of) p) of |
|
214 |
EQUAL => string_ord (pairself Context.theory_name p) |
|
215 |
| order => order |
|
216 |
||
217 |
val thm_ord = theory_ord o pairself theory_of_thm |
|
218 |
||
48326 | 219 |
val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}] |
220 |
||
48318 | 221 |
fun interesting_terms_types_and_classes ctxt prover term_max_depth |
222 |
type_max_depth ts = |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
223 |
let |
48318 | 224 |
fun is_bad_const (x as (s, _)) args = |
225 |
member (op =) atp_logical_consts s orelse |
|
226 |
fst (is_built_in_const_for_prover ctxt prover x args) |
|
48304 | 227 |
fun add_classes @{sort type} = I |
228 |
| add_classes S = union (op =) (map class_name_of S) |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
229 |
fun do_add_type (Type (s, Ts)) = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
230 |
(not (member (op =) bad_types s) ? insert (op =) (type_name_of s)) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
231 |
#> fold do_add_type Ts |
48304 | 232 |
| do_add_type (TFree (_, S)) = add_classes S |
233 |
| do_add_type (TVar (_, S)) = add_classes S |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
234 |
fun add_type T = type_max_depth >= 0 ? do_add_type T |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
235 |
fun mk_app s args = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
236 |
if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")" |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
237 |
else s |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
238 |
fun patternify ~1 _ = "" |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
239 |
| patternify depth t = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
240 |
case strip_comb t of |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
241 |
(Const (s, _), args) => |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
242 |
mk_app (const_name_of s) (map (patternify (depth - 1)) args) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
243 |
| _ => "" |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
244 |
fun add_term_patterns ~1 _ = I |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
245 |
| add_term_patterns depth t = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
246 |
insert (op =) (patternify depth t) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
247 |
#> add_term_patterns (depth - 1) t |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
248 |
val add_term = add_term_patterns term_max_depth |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
249 |
fun add_patterns t = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
250 |
let val (head, args) = strip_comb t in |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
251 |
(case head of |
48318 | 252 |
Const (x as (_, T)) => |
253 |
not (is_bad_const x args) ? (add_term t #> add_type T) |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
254 |
| Free (_, T) => add_type T |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
255 |
| Var (_, T) => add_type T |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
256 |
| Abs (_, T, body) => add_type T #> add_patterns body |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
257 |
| _ => I) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
258 |
#> fold add_patterns args |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
259 |
end |
48326 | 260 |
in [] |> fold add_patterns ts end |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
261 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
262 |
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1}) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
263 |
|
48297 | 264 |
val term_max_depth = 1 |
265 |
val type_max_depth = 1 |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
266 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
267 |
(* TODO: Generate type classes for types? *) |
48318 | 268 |
fun features_of ctxt prover thy status ts = |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
269 |
thy_feature_name_of (Context.theory_name thy) :: |
48318 | 270 |
interesting_terms_types_and_classes ctxt prover term_max_depth type_max_depth |
271 |
ts |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
272 |
|> forall is_lambda_free ts ? cons "no_lams" |
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
273 |
|> forall (not o exists_Const is_exists) ts ? cons "no_skos" |
48302 | 274 |
|> (case status of |
275 |
General => I |
|
276 |
| Induction => cons "induction" |
|
277 |
| Intro => cons "intro" |
|
278 |
| Inductive => cons "inductive" |
|
279 |
| Elim => cons "elim" |
|
280 |
| Simp => cons "simp" |
|
281 |
| Def => cons "def") |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
282 |
|
48326 | 283 |
fun isabelle_dependencies_of all_facts = thms_in_proof (SOME all_facts) |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
284 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
285 |
val freezeT = Type.legacy_freeze_type |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
286 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
287 |
fun freeze (t $ u) = freeze t $ freeze u |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
288 |
| freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
289 |
| freeze (Var ((s, _), T)) = Free (s, freezeT T) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
290 |
| freeze (Const (s, T)) = Const (s, freezeT T) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
291 |
| freeze (Free (s, T)) = Free (s, freezeT T) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
292 |
| freeze t = t |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
293 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
294 |
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
295 |
|
48321 | 296 |
fun run_prover_for_mash ctxt params prover facts goal = |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
297 |
let |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
298 |
val problem = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
299 |
{state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1, |
48289 | 300 |
facts = facts |> map (apfst (apfst (fn name => name ()))) |
48318 | 301 |
|> map Untranslated_Fact} |
48321 | 302 |
val prover = get_minimizing_prover ctxt Normal (K ()) prover |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
303 |
in prover params (K (K (K ""))) problem end |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
304 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
305 |
|
48302 | 306 |
(*** Low-level communication with MaSh ***) |
307 |
||
48323 | 308 |
fun write_file (xs, f) file = |
48318 | 309 |
let val path = Path.explode file in |
48323 | 310 |
File.write path ""; |
311 |
xs |> chunk_list 500 |
|
312 |
|> List.app (File.append path o space_implode "" o map f) |
|
48318 | 313 |
end |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
314 |
|
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
315 |
fun mash_info overlord = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
316 |
if overlord then (getenv "ISABELLE_HOME_USER", "") |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
317 |
else (getenv "ISABELLE_TMP", serial_string ()) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
318 |
|
48382
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
319 |
fun and_rm_files overlord flags files = |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
320 |
if overlord then |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
321 |
"" |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
322 |
else |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
323 |
" && rm -f" ^ flags ^ " -- " ^ |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
324 |
space_implode " " (map File.shell_quote files) |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
325 |
|
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
326 |
fun run_mash ctxt overlord (temp_dir, serial) async core = |
48311 | 327 |
let |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
328 |
val log_file = temp_dir ^ "/mash_log" ^ serial |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
329 |
val err_file = temp_dir ^ "/mash_err" ^ serial |
48311 | 330 |
val command = |
48382
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
331 |
"(" ^ mash_home () ^ "/mash --quiet --outputDir " ^ mash_state_dir () ^ |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
332 |
" --log " ^ log_file ^ " " ^ core ^ ") 2>&1 > " ^ err_file ^ |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
333 |
and_rm_files overlord "" [log_file, err_file] ^ |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
334 |
(if async then " &" else "") |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
335 |
in |
48382
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
336 |
trace_msg ctxt (fn () => |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
337 |
(if async then "Launching " else "Running ") ^ command); |
48323 | 338 |
write_file ([], K "") log_file; |
339 |
write_file ([], K "") err_file; |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
340 |
Isabelle_System.bash command; |
48382
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
341 |
if not async then trace_msg ctxt (K "Done") else () |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
342 |
end |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
343 |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
344 |
(* TODO: Eliminate code once "mash.py" handles sequences of ADD commands as fast |
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
345 |
as a single INIT. *) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
346 |
fun run_mash_init ctxt overlord write_access write_feats write_deps = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
347 |
let |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
348 |
val info as (temp_dir, serial) = mash_info overlord |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
349 |
val in_dir = temp_dir ^ "/mash_init" ^ serial |
48382
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
350 |
|> tap (Path.explode #> Isabelle_System.mkdir) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
351 |
in |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
352 |
write_file write_access (in_dir ^ "/mash_accessibility"); |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
353 |
write_file write_feats (in_dir ^ "/mash_features"); |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
354 |
write_file write_deps (in_dir ^ "/mash_dependencies"); |
48383 | 355 |
run_mash ctxt overlord info false |
48382
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
356 |
("--init --inputDir " ^ in_dir ^ |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
357 |
and_rm_files overlord " -r" [in_dir]) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
358 |
end |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
359 |
|
48383 | 360 |
fun run_mash_commands ctxt overlord save max_suggs write_cmds read_suggs = |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
361 |
let |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
362 |
val info as (temp_dir, serial) = mash_info overlord |
48318 | 363 |
val sugg_file = temp_dir ^ "/mash_suggs" ^ serial |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
364 |
val cmd_file = temp_dir ^ "/mash_commands" ^ serial |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
365 |
in |
48323 | 366 |
write_file ([], K "") sugg_file; |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
367 |
write_file write_cmds cmd_file; |
48383 | 368 |
run_mash ctxt overlord info false |
48318 | 369 |
("--inputFile " ^ cmd_file ^ " --predictions " ^ sugg_file ^ |
370 |
" --numberOfPredictions " ^ string_of_int max_suggs ^ |
|
48382
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
371 |
(if save then " --saveModel" else "") ^ |
641af72b0059
added possibility of running external MaSh commands asynchronously
blanchet
parents:
48381
diff
changeset
|
372 |
(and_rm_files overlord "" [sugg_file, cmd_file])); |
48318 | 373 |
read_suggs (fn () => File.read_lines (Path.explode sugg_file)) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
374 |
end |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
375 |
|
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
376 |
fun str_of_update (name, parents, feats, deps) = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
377 |
"! " ^ escape_meta name ^ ": " ^ escape_metas parents ^ "; " ^ |
48311 | 378 |
escape_metas feats ^ "; " ^ escape_metas deps ^ "\n" |
379 |
||
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
380 |
fun str_of_query (parents, feats) = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
381 |
"? " ^ escape_metas parents ^ "; " ^ escape_metas feats |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
382 |
|
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
383 |
fun init_str_of_update get (upd as (name, _, _, _)) = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
384 |
escape_meta name ^ ": " ^ escape_metas (get upd) ^ "\n" |
48311 | 385 |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
386 |
fun mash_CLEAR ctxt = |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
387 |
let val path = mash_state_dir () |> Path.explode in |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
388 |
trace_msg ctxt (K "MaSh CLEAR"); |
48309 | 389 |
File.fold_dir (fn file => fn () => |
390 |
File.rm (Path.append path (Path.basic file))) |
|
391 |
path () |
|
392 |
end |
|
48302 | 393 |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
394 |
fun mash_INIT ctxt _ [] = mash_CLEAR ctxt |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
395 |
| mash_INIT ctxt overlord upds = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
396 |
(trace_msg ctxt (fn () => "MaSh INIT " ^ |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
397 |
elide_string 1000 (space_implode " " (map #1 upds))); |
48323 | 398 |
run_mash_init ctxt overlord (upds, init_str_of_update #2) |
399 |
(upds, init_str_of_update #3) (upds, init_str_of_update #4)) |
|
48302 | 400 |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
401 |
fun mash_ADD _ _ [] = () |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
402 |
| mash_ADD ctxt overlord upds = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
403 |
(trace_msg ctxt (fn () => "MaSh ADD " ^ |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
404 |
elide_string 1000 (space_implode " " (map #1 upds))); |
48383 | 405 |
run_mash_commands ctxt overlord true 0 (upds, str_of_update) (K ())) |
48302 | 406 |
|
48318 | 407 |
fun mash_QUERY ctxt overlord max_suggs (query as (_, feats)) = |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
408 |
(trace_msg ctxt (fn () => "MaSh QUERY " ^ space_implode " " feats); |
48383 | 409 |
run_mash_commands ctxt overlord false max_suggs |
48323 | 410 |
([query], str_of_query) |
48318 | 411 |
(fn suggs => snd (extract_query (List.last (suggs ())))) |
48311 | 412 |
handle List.Empty => []) |
48302 | 413 |
|
414 |
||
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
415 |
(*** High-level communication with MaSh ***) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
416 |
|
48321 | 417 |
fun try_graph ctxt when def f = |
418 |
f () |
|
419 |
handle Graph.CYCLES (cycle :: _) => |
|
420 |
(trace_msg ctxt (fn () => |
|
421 |
"Cycle involving " ^ commas cycle ^ " when " ^ when); def) |
|
422 |
| Graph.UNDEF name => |
|
423 |
(trace_msg ctxt (fn () => |
|
424 |
"Unknown fact " ^ quote name ^ " when " ^ when); def) |
|
425 |
||
48301 | 426 |
type mash_state = |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
427 |
{thys : bool Symtab.table, |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
428 |
fact_graph : unit Graph.T} |
48301 | 429 |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
430 |
val empty_state = {thys = Symtab.empty, fact_graph = Graph.empty} |
48301 | 431 |
|
432 |
local |
|
433 |
||
48321 | 434 |
fun mash_load _ (state as (true, _)) = state |
435 |
| mash_load ctxt _ = |
|
48309 | 436 |
let val path = mash_state_path () in |
48302 | 437 |
(true, |
438 |
case try File.read_lines path of |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
439 |
SOME (comp_thys :: incomp_thys :: fact_lines) => |
48302 | 440 |
let |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
441 |
fun add_thy comp thy = Symtab.update (thy, comp) |
48322
8a8d71e34297
fixed MaSh state load code so it works even if the facts are read in disorder
blanchet
parents:
48321
diff
changeset
|
442 |
fun add_edge_to name parent = |
8a8d71e34297
fixed MaSh state load code so it works even if the facts are read in disorder
blanchet
parents:
48321
diff
changeset
|
443 |
Graph.default_node (parent, ()) |
8a8d71e34297
fixed MaSh state load code so it works even if the facts are read in disorder
blanchet
parents:
48321
diff
changeset
|
444 |
#> Graph.add_edge (parent, name) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
445 |
fun add_fact_line line = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
446 |
case extract_query line of |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
447 |
("", _) => I (* shouldn't happen *) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
448 |
| (name, parents) => |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
449 |
Graph.default_node (name, ()) |
48322
8a8d71e34297
fixed MaSh state load code so it works even if the facts are read in disorder
blanchet
parents:
48321
diff
changeset
|
450 |
#> fold (add_edge_to name) parents |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
451 |
val thys = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
452 |
Symtab.empty |> fold (add_thy true) (unescape_metas comp_thys) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
453 |
|> fold (add_thy false) (unescape_metas incomp_thys) |
48321 | 454 |
val fact_graph = |
48322
8a8d71e34297
fixed MaSh state load code so it works even if the facts are read in disorder
blanchet
parents:
48321
diff
changeset
|
455 |
try_graph ctxt "loading state" Graph.empty (fn () => |
48321 | 456 |
Graph.empty |> fold add_fact_line fact_lines) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
457 |
in {thys = thys, fact_graph = fact_graph} end |
48304 | 458 |
| _ => empty_state) |
48302 | 459 |
end |
48249 | 460 |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
461 |
fun mash_save ({thys, fact_graph, ...} : mash_state) = |
48301 | 462 |
let |
48309 | 463 |
val path = mash_state_path () |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
464 |
val thys = Symtab.dest thys |
48318 | 465 |
val line_for_thys = escape_metas o AList.find (op =) thys |
466 |
fun fact_line_for name parents = |
|
467 |
escape_meta name ^ ": " ^ escape_metas parents |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
468 |
val append_fact = File.append path o suffix "\n" oo fact_line_for |
48301 | 469 |
in |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
470 |
File.write path (line_for_thys true ^ "\n" ^ line_for_thys false ^ "\n"); |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
471 |
Graph.fold (fn (name, ((), (parents, _))) => fn () => |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
472 |
append_fact name (Graph.Keys.dest parents)) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
473 |
fact_graph () |
48301 | 474 |
end |
475 |
||
48302 | 476 |
val global_state = |
48381 | 477 |
Synchronized.var "Sledgehammer_MaSh.global_state" (false, empty_state) |
48301 | 478 |
|
479 |
in |
|
480 |
||
48321 | 481 |
fun mash_map ctxt f = |
482 |
Synchronized.change global_state (mash_load ctxt ##> (f #> tap mash_save)) |
|
48302 | 483 |
|
48321 | 484 |
fun mash_get ctxt = |
485 |
Synchronized.change_result global_state (mash_load ctxt #> `snd) |
|
48302 | 486 |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
487 |
fun mash_unlearn ctxt = |
48302 | 488 |
Synchronized.change global_state (fn _ => |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
489 |
(mash_CLEAR ctxt; File.write (mash_state_path ()) ""; |
48308 | 490 |
(true, empty_state))) |
48301 | 491 |
|
492 |
end |
|
493 |
||
48318 | 494 |
fun mash_could_suggest_facts () = mash_home () <> "" |
48321 | 495 |
fun mash_can_suggest_facts ctxt = |
496 |
not (Graph.is_empty (#fact_graph (mash_get ctxt))) |
|
48249 | 497 |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
498 |
fun parents_wrt_facts facts fact_graph = |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
499 |
let |
48378 | 500 |
val facts = [] |> fold (cons o nickname_of o snd) facts |
48330 | 501 |
val tab = Symtab.empty |> fold (fn name => Symtab.update (name, ())) facts |
48377 | 502 |
fun insert_not_seen seen name = |
503 |
not (member (op =) seen name) ? insert (op =) name |
|
48378 | 504 |
fun parents_of _ parents [] = parents |
48377 | 505 |
| parents_of seen parents (name :: names) = |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
506 |
if Symtab.defined tab name then |
48377 | 507 |
parents_of (name :: seen) (name :: parents) names |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
508 |
else |
48377 | 509 |
parents_of (name :: seen) parents |
510 |
(Graph.Keys.fold (insert_not_seen seen) |
|
511 |
(Graph.imm_preds fact_graph name) names) |
|
512 |
in parents_of [] [] (Graph.maximals fact_graph) end |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
513 |
|
48318 | 514 |
(* Generate more suggestions than requested, because some might be thrown out |
515 |
later for various reasons and "meshing" gives better results with some |
|
516 |
slack. *) |
|
517 |
fun max_suggs_of max_facts = max_facts + Int.min (200, max_facts) |
|
518 |
||
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
519 |
fun is_fact_in_graph fact_graph (_, th) = |
48378 | 520 |
can (Graph.get_node fact_graph) (nickname_of th) |
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
521 |
|
48318 | 522 |
fun mash_suggest_facts ctxt ({overlord, ...} : params) prover max_facts hyp_ts |
523 |
concl_t facts = |
|
48301 | 524 |
let |
48302 | 525 |
val thy = Proof_Context.theory_of ctxt |
48321 | 526 |
val fact_graph = #fact_graph (mash_get ctxt) |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
527 |
val parents = parents_wrt_facts facts fact_graph |
48318 | 528 |
val feats = features_of ctxt prover thy General (concl_t :: hyp_ts) |
529 |
val suggs = |
|
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
530 |
if Graph.is_empty fact_graph then [] |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
531 |
else mash_QUERY ctxt overlord (max_suggs_of max_facts) (parents, feats) |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
532 |
val selected = facts |> suggested_facts suggs |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
533 |
val unknown = facts |> filter_out (is_fact_in_graph fact_graph) |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
534 |
in (selected, unknown) end |
48249 | 535 |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
536 |
fun add_thys_for thy = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
537 |
let fun add comp thy = Symtab.update (Context.theory_name thy, comp) in |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
538 |
add false thy #> fold (add true) (Theory.ancestors_of thy) |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
539 |
end |
48249 | 540 |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
541 |
fun update_fact_graph ctxt (name, parents, feats, deps) (upds, graph) = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
542 |
let |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
543 |
fun maybe_add_from from (accum as (parents, graph)) = |
48321 | 544 |
try_graph ctxt "updating graph" accum (fn () => |
545 |
(from :: parents, Graph.add_edge_acyclic (from, name) graph)) |
|
546 |
val graph = graph |> Graph.default_node (name, ()) |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
547 |
val (parents, graph) = ([], graph) |> fold maybe_add_from parents |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
548 |
val (deps, graph) = ([], graph) |> fold maybe_add_from deps |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
549 |
in ((name, parents, feats, deps) :: upds, graph) end |
48306 | 550 |
|
48384
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
551 |
val learn_timeout_slack = 2.0 |
48318 | 552 |
|
48384
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
553 |
fun launch_thread timeout task = |
48383 | 554 |
let |
48384
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
555 |
val hard_timeout = time_mult learn_timeout_slack timeout |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
556 |
val birth_time = Time.now () |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
557 |
val death_time = Time.+ (birth_time, hard_timeout) |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
558 |
val desc = ("machine learner for Sledgehammer", "") |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
559 |
in Async_Manager.launch MaShN birth_time death_time desc task end |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
560 |
|
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
561 |
fun mash_learn_proof ctxt ({overlord, timeout, ...} : params) prover t facts |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
562 |
used_ths = |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
563 |
if is_smt_prover ctxt prover then |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
564 |
() |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
565 |
else |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
566 |
launch_thread timeout |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
567 |
(fn () => |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
568 |
let |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
569 |
val thy = Proof_Context.theory_of ctxt |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
570 |
val name = timestamp () ^ " " ^ serial_string () (* freshish *) |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
571 |
val feats = features_of ctxt prover thy General [t] |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
572 |
val deps = used_ths |> map nickname_of |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
573 |
in |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
574 |
mash_map ctxt (fn {thys, fact_graph} => |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
575 |
let |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
576 |
val parents = parents_wrt_facts facts fact_graph |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
577 |
val upds = [(name, parents, feats, deps)] |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
578 |
val (upds, fact_graph) = |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
579 |
([], fact_graph) |> fold (update_fact_graph ctxt) upds |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
580 |
in |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
581 |
mash_ADD ctxt overlord upds; |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
582 |
{thys = thys, fact_graph = fact_graph} |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
583 |
end); |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
584 |
(true, "") |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
585 |
end) |
48383 | 586 |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
587 |
(* Too many dependencies is a sign that a decision procedure is at work. There |
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
588 |
isn't much too learn from such proofs. *) |
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
589 |
val max_dependencies = 10 |
48384
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
590 |
val pass1_learn_timeout_factor = 0.75 |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
591 |
|
48318 | 592 |
(* The timeout is understood in a very slack fashion. *) |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
593 |
fun mash_learn_thy ctxt ({provers, verbose, overlord, ...} : params) thy timeout |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
594 |
facts = |
48304 | 595 |
let |
48318 | 596 |
val timer = Timer.startRealTimer () |
597 |
val prover = hd provers |
|
598 |
fun timed_out frac = |
|
599 |
Time.> (Timer.checkRealTimer timer, time_mult frac timeout) |
|
48321 | 600 |
val {fact_graph, ...} = mash_get ctxt |
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
601 |
val new_facts = |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
602 |
facts |> filter_out (is_fact_in_graph fact_graph) |
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
603 |
|> sort (thm_ord o pairself snd) |
48308 | 604 |
in |
605 |
if null new_facts then |
|
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
606 |
"" |
48308 | 607 |
else |
48304 | 608 |
let |
48308 | 609 |
val ths = facts |> map snd |
48315
82d6e46c673f
fixed order of accessibles + other tweaks to MaSh
blanchet
parents:
48314
diff
changeset
|
610 |
val all_names = |
48324 | 611 |
ths |> filter_out is_likely_tautology_or_too_meta |
48378 | 612 |
|> map (rpair () o nickname_of) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
613 |
|> Symtab.make |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
614 |
fun trim_deps deps = if length deps > max_dependencies then [] else deps |
48318 | 615 |
fun do_fact _ (accum as (_, true)) = accum |
616 |
| do_fact ((_, (_, status)), th) ((parents, upds), false) = |
|
617 |
let |
|
48378 | 618 |
val name = nickname_of th |
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
619 |
val feats = |
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
620 |
features_of ctxt prover (theory_of_thm th) status [prop_of th] |
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
621 |
val deps = isabelle_dependencies_of all_names th |> trim_deps |
48318 | 622 |
val upd = (name, parents, feats, deps) |
623 |
in (([name], upd :: upds), timed_out pass1_learn_timeout_factor) end |
|
48332
271a4a6af734
optimize parent computation in MaSh + remove temporary files
blanchet
parents:
48330
diff
changeset
|
624 |
val parents = parents_wrt_facts facts fact_graph |
48318 | 625 |
val ((_, upds), _) = |
626 |
((parents, []), false) |> fold do_fact new_facts |>> apsnd rev |
|
627 |
val n = length upds |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
628 |
fun trans {thys, fact_graph} = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
629 |
let |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
630 |
val mash_INIT_or_ADD = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
631 |
if Graph.is_empty fact_graph then mash_INIT else mash_ADD |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
632 |
val (upds, fact_graph) = |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
633 |
([], fact_graph) |> fold (update_fact_graph ctxt) upds |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
634 |
in |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
635 |
(mash_INIT_or_ADD ctxt overlord (rev upds); |
48383 | 636 |
{thys = thys |> add_thys_for thy, fact_graph = fact_graph}) |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
637 |
end |
48318 | 638 |
in |
48321 | 639 |
mash_map ctxt trans; |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
640 |
if verbose then |
48383 | 641 |
"Learned " ^ string_of_int n ^ " proof" ^ plural_s n ^ |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
642 |
(if verbose then |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
643 |
" in " ^ string_from_time (Timer.checkRealTimer timer) |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
644 |
else |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
645 |
"") ^ "." |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
646 |
else |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
647 |
"" |
48318 | 648 |
end |
48308 | 649 |
end |
48304 | 650 |
|
48383 | 651 |
fun mash_learn ctxt params = |
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
652 |
let |
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
653 |
val thy = Proof_Context.theory_of ctxt |
48383 | 654 |
val _ = Output.urgent_message "MaShing..." |
655 |
val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt |
|
656 |
val facts = all_facts_of thy css_table |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
657 |
in |
48383 | 658 |
mash_learn_thy ctxt params thy infinite_timeout facts |
659 |
|> (fn "" => "Learned nothing." | msg => msg) |
|
660 |
|> Output.urgent_message |
|
48316
252f45c04042
drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents:
48315
diff
changeset
|
661 |
end |
48249 | 662 |
|
48318 | 663 |
(* The threshold should be large enough so that MaSh doesn't kick in for Auto |
664 |
Sledgehammer and Try. *) |
|
665 |
val min_secs_for_learning = 15 |
|
666 |
||
48321 | 667 |
fun relevant_facts ctxt (params as {learn, fact_filter, timeout, ...}) prover |
668 |
max_facts ({add, only, ...} : fact_override) hyp_ts concl_t facts = |
|
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
669 |
if not (subset (op =) (the_list fact_filter, fact_filters)) then |
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
670 |
error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".") |
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
671 |
else if only then |
48289 | 672 |
facts |
48321 | 673 |
else if max_facts <= 0 orelse null facts then |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
674 |
[] |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
675 |
else |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
676 |
let |
48318 | 677 |
val thy = Proof_Context.theory_of ctxt |
48327
568b3193e53e
don't include hidden facts in relevance filter + tweak MaSh learning
blanchet
parents:
48326
diff
changeset
|
678 |
fun maybe_learn () = |
48384
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
679 |
if learn andalso not (Async_Manager.has_running_threads MaShN) andalso |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
680 |
Time.toSeconds timeout >= min_secs_for_learning then |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
681 |
let val timeout = time_mult learn_timeout_slack timeout in |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
682 |
launch_thread timeout |
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents:
48383
diff
changeset
|
683 |
(fn () => (true, mash_learn_thy ctxt params thy timeout facts)) |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
684 |
end |
48318 | 685 |
else |
686 |
() |
|
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
687 |
val fact_filter = |
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
688 |
case fact_filter of |
48379
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
689 |
SOME ff => (() |> ff <> mepoN ? maybe_learn; ff) |
48318 | 690 |
| NONE => |
48327
568b3193e53e
don't include hidden facts in relevance filter + tweak MaSh learning
blanchet
parents:
48326
diff
changeset
|
691 |
if mash_can_suggest_facts ctxt then (maybe_learn (); meshN) |
48379
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
692 |
else if mash_could_suggest_facts () then (maybe_learn (); mepoN) |
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
693 |
else mepoN |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
694 |
val add_ths = Attrib.eval_thms ctxt add |
48292 | 695 |
fun prepend_facts ths accepts = |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
696 |
((facts |> filter (member Thm.eq_thm_prop ths o snd)) @ |
48292 | 697 |
(accepts |> filter_out (member Thm.eq_thm_prop ths o snd))) |
48293 | 698 |
|> take max_facts |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
699 |
fun iter () = |
48298 | 700 |
iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts |
701 |
concl_t facts |
|
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
702 |
fun mash () = |
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
703 |
mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts |
48314
ee33ba3c0e05
added option to control which fact filter is used
blanchet
parents:
48313
diff
changeset
|
704 |
val mess = |
48320
891a24a48155
improved meshing of MaSh and Meng--Paulson if some MaSh suggestions are cut-off (the common case)
blanchet
parents:
48319
diff
changeset
|
705 |
[] |> (if fact_filter <> mashN then cons (iter (), []) else I) |
48379
2b5ad61e2ccc
renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents:
48378
diff
changeset
|
706 |
|> (if fact_filter <> mepoN then cons (mash ()) else I) |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
707 |
in |
48313
0faafdffa662
mesh facts by taking into consideration whether a fact is known to MeSh
blanchet
parents:
48312
diff
changeset
|
708 |
mesh_facts max_facts mess |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
709 |
|> not (null add_ths) ? prepend_facts add_ths |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
710 |
end |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
711 |
|
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
712 |
fun kill_learners () = Async_Manager.kill_threads MaShN "learner" |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
713 |
fun running_learners () = Async_Manager.running_threads MaShN "learner" |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
48318
diff
changeset
|
714 |
|
48248 | 715 |
end; |