author | blanchet |
Wed, 18 Jul 2012 08:44:04 +0200 | |
changeset 48312 | b40722a81ac9 |
parent 48311 | 3c4e10606567 |
child 48313 | 0faafdffa662 |
permissions | -rw-r--r-- |
48248 | 1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML |
2 |
Author: Jasmin Blanchette, TU Muenchen |
|
3 |
||
4 |
Sledgehammer's machine-learning-based relevance filter (MaSh). |
|
5 |
*) |
|
6 |
||
7 |
signature SLEDGEHAMMER_FILTER_MASH = |
|
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 |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
18 |
val escape_meta : string -> string |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
19 |
val escape_metas : string list -> string |
48308 | 20 |
val unescape_meta : string -> string |
21 |
val unescape_metas : string -> string list |
|
48311 | 22 |
val extract_query : string -> string * string list |
48312 | 23 |
val suggested_facts : string list -> fact list -> fact list |
24 |
val mesh_facts : int -> fact list -> fact list -> fact list -> fact list |
|
48299 | 25 |
val all_non_tautological_facts_of : |
26 |
theory -> status Termtab.table -> fact list |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
27 |
val theory_ord : theory * theory -> order |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
28 |
val thm_ord : thm * thm -> order |
48312 | 29 |
val thy_facts_from_thms : fact list -> string list Symtab.table |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
30 |
val has_thy : theory -> thm -> bool |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
31 |
val parent_facts : theory -> string list Symtab.table -> string list |
48302 | 32 |
val features_of : theory -> status -> term list -> string list |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
33 |
val isabelle_dependencies_of : string list -> thm -> string list |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
34 |
val goal_of_thm : theory -> thm -> thm |
48296
e7f01b7e244e
gracefully handle the case of empty theories when going up the accessibility chain
blanchet
parents:
48293
diff
changeset
|
35 |
val run_prover : Proof.context -> params -> fact list -> thm -> prover_result |
48304 | 36 |
val thy_name_of_fact : string -> string |
48308 | 37 |
val mash_RESET : Proof.context -> unit |
38 |
val mash_ADD : |
|
39 |
Proof.context -> (string * string list * string list * string list) list |
|
40 |
-> unit |
|
41 |
val mash_DEL : Proof.context -> string list -> string list -> unit |
|
48311 | 42 |
val mash_QUERY : Proof.context -> string list * string list -> string list |
48308 | 43 |
val mash_reset : Proof.context -> unit |
44 |
val mash_can_suggest_facts : Proof.context -> bool |
|
48298 | 45 |
val mash_suggest_facts : |
48311 | 46 |
Proof.context -> params -> string -> term list -> term -> fact list |
47 |
-> fact list |
|
48308 | 48 |
val mash_can_learn_thy : Proof.context -> theory -> bool |
49 |
val mash_learn_thy : Proof.context -> theory -> real -> unit |
|
48309 | 50 |
val mash_learn_proof : Proof.context -> theory -> term -> thm list -> unit |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
51 |
val relevant_facts : |
48292 | 52 |
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
|
53 |
-> term -> fact list -> fact list |
48248 | 54 |
end; |
55 |
||
56 |
structure Sledgehammer_Filter_MaSh : SLEDGEHAMMER_FILTER_MASH = |
|
57 |
struct |
|
48249 | 58 |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
59 |
open ATP_Util |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
60 |
open ATP_Problem_Generate |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
61 |
open Sledgehammer_Util |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
62 |
open Sledgehammer_Fact |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
63 |
open Sledgehammer_Filter_Iter |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
64 |
open Sledgehammer_Provers |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
65 |
|
48308 | 66 |
val trace = |
67 |
Attrib.setup_config_bool @{binding sledgehammer_filter_mash_trace} (K false) |
|
68 |
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else () |
|
69 |
||
48309 | 70 |
fun mash_dir () = |
71 |
getenv "ISABELLE_HOME_USER" ^ "/mash" |
|
72 |
|> tap (fn dir => Isabelle_System.mkdir (Path.explode dir)) |
|
73 |
fun mash_state_path () = mash_dir () ^ "/state" |> Path.explode |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
74 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
75 |
(*** Isabelle helpers ***) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
76 |
|
48308 | 77 |
fun meta_char c = |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
78 |
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
|
79 |
c = #")" orelse c = #"," then |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
80 |
String.str c |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
81 |
else |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
82 |
(* fixed width, in case more digits follow *) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
83 |
"\\" ^ stringN_of_int 3 (Char.ord c) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
84 |
|
48308 | 85 |
fun unmeta_chars accum [] = String.implode (rev accum) |
86 |
| unmeta_chars accum (#"\\" :: d1 :: d2 :: d3 :: cs) = |
|
87 |
(case Int.fromString (String.implode [d1, d2, d3]) of |
|
88 |
SOME n => unmeta_chars (Char.chr n :: accum) cs |
|
89 |
| NONE => "" (* error *)) |
|
90 |
| unmeta_chars _ (#"\\" :: _) = "" (* error *) |
|
91 |
| unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs |
|
92 |
||
93 |
val escape_meta = String.translate meta_char |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
94 |
val escape_metas = map escape_meta #> space_implode " " |
48308 | 95 |
val unescape_meta = unmeta_chars [] o String.explode |
96 |
val unescape_metas = map unescape_meta o space_explode " " |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
97 |
|
48311 | 98 |
val explode_suggs = |
99 |
space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta |
|
100 |
fun extract_query line = |
|
101 |
case space_explode ":" line of |
|
102 |
[goal_name, suggs] => (unescape_meta goal_name, explode_suggs suggs) |
|
48312 | 103 |
| _ => ("", []) |
48311 | 104 |
|
105 |
fun find_suggested facts sugg = |
|
106 |
find_first (fn (_, th) => Thm.get_name_hint th = sugg) facts |
|
107 |
fun suggested_facts suggs facts = map_filter (find_suggested facts) suggs |
|
108 |
||
48312 | 109 |
fun mesh_facts max_facts iter_facts mash_facts mash_rejects = |
110 |
let |
|
111 |
val fact_eq = Thm.eq_thm o pairself snd |
|
112 |
val num_iter = length iter_facts |
|
113 |
val num_mash = length mash_facts |
|
114 |
fun score_in f fs n = |
|
115 |
case find_index (curry fact_eq f) fs of |
|
116 |
~1 => length fs |
|
117 |
| j => j |
|
118 |
fun score_of fact = |
|
119 |
score_in fact iter_facts num_iter + score_in fact mash_facts num_mash |
|
120 |
in |
|
121 |
union fact_eq iter_facts mash_facts |
|
122 |
|> map (`score_of) |> sort (int_ord o pairself fst) |> map snd |
|
123 |
|> take max_facts |
|
124 |
end |
|
125 |
||
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
126 |
val thy_feature_prefix = "y_" |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
127 |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
128 |
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
|
129 |
val const_name_of = prefix const_prefix |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
130 |
val type_name_of = prefix type_const_prefix |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
131 |
val class_name_of = prefix class_prefix |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
132 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
133 |
local |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
134 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
135 |
fun has_bool @{typ bool} = true |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
136 |
| has_bool (Type (_, Ts)) = exists has_bool Ts |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
137 |
| has_bool _ = false |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
138 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
139 |
fun has_fun (Type (@{type_name fun}, _)) = true |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
140 |
| has_fun (Type (_, Ts)) = exists has_fun Ts |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
141 |
| has_fun _ = false |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
142 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
143 |
val is_conn = member (op =) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
144 |
[@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj}, |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
145 |
@{const_name HOL.implies}, @{const_name Not}, |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
146 |
@{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex}, |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
147 |
@{const_name HOL.eq}] |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
148 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
149 |
val has_bool_arg_const = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
150 |
exists_Const (fn (c, T) => |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
151 |
not (is_conn c) andalso exists has_bool (binder_types T)) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
152 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
153 |
fun higher_inst_const thy (c, T) = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
154 |
case binder_types T of |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
155 |
[] => false |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
156 |
| Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
157 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
158 |
val binders = [@{const_name All}, @{const_name Ex}] |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
159 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
160 |
in |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
161 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
162 |
fun is_fo_term thy t = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
163 |
let |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
164 |
val t = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
165 |
t |> Envir.beta_eta_contract |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
166 |
|> transform_elim_prop |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
167 |
|> Object_Logic.atomize_term thy |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
168 |
in |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
169 |
Term.is_first_order binders t andalso |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
170 |
not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
171 |
| _ => false) t orelse |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
172 |
has_bool_arg_const t orelse exists_Const (higher_inst_const thy) t) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
173 |
end |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
174 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
175 |
end |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
176 |
|
48302 | 177 |
fun interesting_terms_types_and_classes term_max_depth type_max_depth ts = |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
178 |
let |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
179 |
val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}] |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
180 |
val bad_consts = atp_widely_irrelevant_consts |
48304 | 181 |
fun add_classes @{sort type} = I |
182 |
| 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
|
183 |
fun do_add_type (Type (s, Ts)) = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
184 |
(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
|
185 |
#> fold do_add_type Ts |
48304 | 186 |
| do_add_type (TFree (_, S)) = add_classes S |
187 |
| do_add_type (TVar (_, S)) = add_classes S |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
188 |
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
|
189 |
fun mk_app s args = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
190 |
if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")" |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
191 |
else s |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
192 |
fun patternify ~1 _ = "" |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
193 |
| patternify depth t = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
194 |
case strip_comb t of |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
195 |
(Const (s, _), args) => |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
196 |
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
|
197 |
| _ => "" |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
198 |
fun add_term_patterns ~1 _ = I |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
199 |
| add_term_patterns depth t = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
200 |
insert (op =) (patternify depth t) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
201 |
#> add_term_patterns (depth - 1) t |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
202 |
val add_term = add_term_patterns term_max_depth |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
203 |
fun add_patterns t = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
204 |
let val (head, args) = strip_comb t in |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
205 |
(case head of |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
206 |
Const (s, T) => |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
207 |
not (member (op =) bad_consts s) ? (add_term t #> add_type T) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
208 |
| Free (_, T) => add_type T |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
209 |
| Var (_, T) => add_type T |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
210 |
| Abs (_, T, body) => add_type T #> add_patterns body |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
211 |
| _ => I) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
212 |
#> fold add_patterns args |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
213 |
end |
48302 | 214 |
in [] |> fold add_patterns ts |> sort string_ord end |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
215 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
216 |
fun is_likely_tautology th = |
48302 | 217 |
null (interesting_terms_types_and_classes 0 ~1 [prop_of th]) andalso |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
218 |
not (Thm.eq_thm_prop (@{thm ext}, th)) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
219 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
220 |
fun is_too_meta thy th = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
221 |
fastype_of (Object_Logic.atomize_term thy (prop_of th)) <> @{typ bool} |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
222 |
|
48299 | 223 |
fun all_non_tautological_facts_of thy css_table = |
224 |
all_facts_of thy css_table |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
225 |
|> filter_out ((is_likely_tautology orf is_too_meta thy) o snd) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
226 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
227 |
fun theory_ord p = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
228 |
if Theory.eq_thy p then EQUAL |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
229 |
else if Theory.subthy p then LESS |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
230 |
else if Theory.subthy (swap p) then GREATER |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
231 |
else EQUAL |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
232 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
233 |
val thm_ord = theory_ord o pairself theory_of_thm |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
234 |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
235 |
(* ### FIXME: optimize *) |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
236 |
fun thy_facts_from_thms ths = |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
237 |
ths |> map (snd #> `(theory_of_thm #> Context.theory_name)) |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
238 |
|> AList.group (op =) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
239 |
|> sort (int_ord o pairself (length o Theory.ancestors_of o theory_of_thm |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
240 |
o hd o snd)) |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
241 |
|> map (apsnd (sort (rev_order o thm_ord) #> map Thm.get_name_hint)) |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
242 |
|> Symtab.make |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
243 |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
244 |
fun has_thy thy th = |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
245 |
Context.theory_name thy = Context.theory_name (theory_of_thm th) |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
246 |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
247 |
fun parent_facts thy thy_facts = |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
248 |
let |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
249 |
fun add_last thy = |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
250 |
case Symtab.lookup thy_facts (Context.theory_name thy) of |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
251 |
SOME (last_fact :: _) => insert (op =) last_fact |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
252 |
| _ => add_parent thy |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
253 |
and add_parent thy = fold add_last (Theory.parents_of thy) |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
254 |
in add_parent thy [] end |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
255 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
256 |
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
|
257 |
|
48297 | 258 |
val term_max_depth = 1 |
259 |
val type_max_depth = 1 |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
260 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
261 |
(* TODO: Generate type classes for types? *) |
48302 | 262 |
fun features_of thy status ts = |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
263 |
thy_feature_name_of (Context.theory_name thy) :: |
48302 | 264 |
interesting_terms_types_and_classes term_max_depth type_max_depth ts |
265 |
|> exists (not o is_lambda_free) ts ? cons "lambdas" |
|
266 |
|> exists (exists_Const is_exists) ts ? cons "skolems" |
|
267 |
|> exists (not o is_fo_term thy) ts ? cons "ho" |
|
268 |
|> (case status of |
|
269 |
General => I |
|
270 |
| Induction => cons "induction" |
|
271 |
| Intro => cons "intro" |
|
272 |
| Inductive => cons "inductive" |
|
273 |
| Elim => cons "elim" |
|
274 |
| Simp => cons "simp" |
|
275 |
| Def => cons "def") |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
276 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
277 |
fun isabelle_dependencies_of all_facts = |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
278 |
thms_in_proof (SOME all_facts) #> sort string_ord |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
279 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
280 |
val freezeT = Type.legacy_freeze_type |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
281 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
282 |
fun freeze (t $ u) = freeze t $ freeze u |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
283 |
| 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
|
284 |
| freeze (Var ((s, _), T)) = Free (s, freezeT T) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
285 |
| freeze (Const (s, T)) = Const (s, freezeT T) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
286 |
| freeze (Free (s, T)) = Free (s, freezeT T) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
287 |
| freeze t = t |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
288 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
289 |
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
|
290 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
291 |
fun run_prover ctxt (params as {provers, ...}) facts goal = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
292 |
let |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
293 |
val problem = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
294 |
{state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1, |
48289 | 295 |
facts = facts |> map (apfst (apfst (fn name => name ()))) |
296 |
|> map Sledgehammer_Provers.Untranslated_Fact} |
|
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
297 |
val prover = |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
298 |
Sledgehammer_Minimize.get_minimizing_prover ctxt |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
299 |
Sledgehammer_Provers.Normal (hd provers) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
300 |
in prover params (K (K (K ""))) problem end |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
301 |
|
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
302 |
fun accessibility_of thy thy_facts = |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
303 |
case Symtab.lookup thy_facts (Context.theory_name thy) of |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
304 |
SOME (fact :: _) => [fact] |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
305 |
| _ => parent_facts thy thy_facts |
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
306 |
|
48304 | 307 |
val thy_name_of_fact = hd o Long_Name.explode |
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
308 |
|
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
309 |
|
48302 | 310 |
(*** Low-level communication with MaSh ***) |
311 |
||
48311 | 312 |
fun run_mash ctxt save write_cmds read_preds = |
313 |
let |
|
314 |
val temp_dir = getenv "ISABELLE_TMP" |
|
315 |
val serial = serial_string () |
|
316 |
val cmd_file = temp_dir ^ "/mash_commands." ^ serial |
|
317 |
val cmd_path = Path.explode cmd_file |
|
318 |
val pred_file = temp_dir ^ "/mash_preds." ^ serial |
|
319 |
val log_file = temp_dir ^ "/mash_log." ^ serial |
|
320 |
val command = |
|
321 |
getenv "MASH_HOME" ^ "/mash.py --inputFile " ^ cmd_file ^ |
|
322 |
" --outputDir " ^ mash_dir () ^ " --predictions " ^ pred_file ^ |
|
323 |
" --log " ^ log_file ^ (if save then " --saveModel" else "") ^ |
|
324 |
" > /dev/null" |
|
325 |
val _ = File.write cmd_path "" |
|
326 |
val _ = write_cmds (File.append cmd_path) |
|
327 |
val _ = trace_msg ctxt (fn () => " running " ^ command) |
|
328 |
val _ = Isabelle_System.bash command |
|
329 |
in read_preds (fn () => File.read_lines (Path.explode pred_file)) end |
|
330 |
||
331 |
fun str_of_update (fact, access, feats, deps) = |
|
332 |
"! " ^ escape_meta fact ^ ": " ^ escape_metas access ^ "; " ^ |
|
333 |
escape_metas feats ^ "; " ^ escape_metas deps ^ "\n" |
|
334 |
||
335 |
fun str_of_query (access, feats) = |
|
336 |
"? " ^ escape_metas access ^ "; " ^ escape_metas feats |
|
337 |
||
48308 | 338 |
fun mash_RESET ctxt = |
48309 | 339 |
let val path = mash_dir () |> Path.explode in |
340 |
trace_msg ctxt (K "MaSh RESET"); |
|
341 |
File.fold_dir (fn file => fn () => |
|
342 |
File.rm (Path.append path (Path.basic file))) |
|
343 |
path () |
|
344 |
end |
|
48302 | 345 |
|
48309 | 346 |
fun mash_ADD _ [] = () |
48311 | 347 |
| mash_ADD ctxt upds = |
348 |
(trace_msg ctxt (fn () => "MaSh ADD " ^ space_implode " " (map #1 upds)); |
|
349 |
run_mash ctxt true (fn append => List.app (append o str_of_update) upds) |
|
350 |
(K ())) |
|
48302 | 351 |
|
48308 | 352 |
fun mash_DEL ctxt facts feats = |
353 |
trace_msg ctxt (fn () => |
|
354 |
"MaSh DEL " ^ escape_metas facts ^ "; " ^ escape_metas feats) |
|
48302 | 355 |
|
48311 | 356 |
fun mash_QUERY ctxt (query as (_, feats)) = |
357 |
(trace_msg ctxt (fn () => "MaSh SUGGEST " ^ space_implode " " feats); |
|
358 |
run_mash ctxt false (fn append => append (str_of_query query)) |
|
359 |
(fn preds => snd (extract_query (List.last (preds ())))) |
|
360 |
handle List.Empty => []) |
|
48302 | 361 |
|
362 |
||
48251
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
363 |
(*** High-level communication with MaSh ***) |
6cdcfbddc077
moved most of MaSh exporter code to Sledgehammer
blanchet
parents:
48249
diff
changeset
|
364 |
|
48301 | 365 |
type mash_state = |
48308 | 366 |
{dirty_thys : unit Symtab.table, |
48301 | 367 |
thy_facts : string list Symtab.table} |
368 |
||
48304 | 369 |
val empty_state = |
48308 | 370 |
{dirty_thys = Symtab.empty, |
48301 | 371 |
thy_facts = Symtab.empty} |
372 |
||
373 |
local |
|
374 |
||
375 |
fun mash_load (state as (true, _)) = state |
|
376 |
| mash_load _ = |
|
48309 | 377 |
let val path = mash_state_path () in |
48302 | 378 |
(true, |
379 |
case try File.read_lines path of |
|
48308 | 380 |
SOME (dirty_line :: facts_lines) => |
48302 | 381 |
let |
48306 | 382 |
fun dirty_thys_of_line line = |
48308 | 383 |
Symtab.make (line |> unescape_metas |> map (rpair ())) |
48302 | 384 |
fun add_facts_line line = |
48308 | 385 |
case unescape_metas line of |
48302 | 386 |
thy :: facts => Symtab.update_new (thy, facts) |
387 |
| _ => I (* shouldn't happen *) |
|
388 |
in |
|
48308 | 389 |
{dirty_thys = dirty_thys_of_line dirty_line, |
48302 | 390 |
thy_facts = fold add_facts_line facts_lines Symtab.empty} |
391 |
end |
|
48304 | 392 |
| _ => empty_state) |
48302 | 393 |
end |
48249 | 394 |
|
48308 | 395 |
fun mash_save ({dirty_thys, thy_facts} : mash_state) = |
48301 | 396 |
let |
48309 | 397 |
val path = mash_state_path () |
48308 | 398 |
val dirty_line = (escape_metas (Symtab.keys dirty_thys)) ^ "\n" |
48303
f1d135d0ea69
improved MaSh string escaping and make more operations string-based
blanchet
parents:
48302
diff
changeset
|
399 |
fun fact_line_for (thy, facts) = escape_metas (thy :: facts) ^ "\n" |
48301 | 400 |
in |
48308 | 401 |
File.write path dirty_line; |
48301 | 402 |
Symtab.fold (fn thy_fact => fn () => |
48308 | 403 |
File.append path (fact_line_for thy_fact)) thy_facts () |
48301 | 404 |
end |
405 |
||
48302 | 406 |
val global_state = |
48304 | 407 |
Synchronized.var "Sledgehammer_Filter_MaSh.global_state" (false, empty_state) |
48301 | 408 |
|
409 |
in |
|
410 |
||
48306 | 411 |
fun mash_map f = |
48302 | 412 |
Synchronized.change global_state (mash_load ##> (f #> tap mash_save)) |
413 |
||
48304 | 414 |
fun mash_get () = Synchronized.change_result global_state (mash_load #> `snd) |
48302 | 415 |
|
48308 | 416 |
fun mash_reset ctxt = |
48302 | 417 |
Synchronized.change global_state (fn _ => |
48309 | 418 |
(mash_RESET ctxt; File.write (mash_state_path ()) ""; |
48308 | 419 |
(true, empty_state))) |
48301 | 420 |
|
421 |
end |
|
422 |
||
48308 | 423 |
fun mash_can_suggest_facts (_ : Proof.context) = |
424 |
not (Symtab.is_empty (#thy_facts (mash_get ()))) |
|
48249 | 425 |
|
48311 | 426 |
fun mash_suggest_facts ctxt params prover hyp_ts concl_t facts = |
48301 | 427 |
let |
48302 | 428 |
val thy = Proof_Context.theory_of ctxt |
48311 | 429 |
val thy_facts = #thy_facts (mash_get ()) |
430 |
val access = accessibility_of thy thy_facts |
|
48302 | 431 |
val feats = features_of thy General (concl_t :: hyp_ts) |
48311 | 432 |
val suggs = mash_QUERY ctxt (access, feats) |
433 |
in suggested_facts suggs facts end |
|
48249 | 434 |
|
48308 | 435 |
fun mash_can_learn_thy (_ : Proof.context) thy = |
48306 | 436 |
not (Symtab.defined (#dirty_thys (mash_get ())) (Context.theory_name thy)) |
48304 | 437 |
|
438 |
fun is_fact_in_thy_facts thy_facts fact = |
|
439 |
case Symtab.lookup thy_facts (thy_name_of_fact fact) of |
|
440 |
SOME facts => member (op =) facts fact |
|
441 |
| NONE => false |
|
48249 | 442 |
|
48306 | 443 |
fun zip_facts news [] = news |
444 |
| zip_facts [] olds = olds |
|
445 |
| zip_facts (new :: news) (old :: olds) = |
|
446 |
if new = old then |
|
447 |
new :: zip_facts news olds |
|
448 |
else if member (op =) news old then |
|
449 |
old :: zip_facts (filter_out (curry (op =) old) news) olds |
|
450 |
else if member (op =) olds new then |
|
451 |
new :: zip_facts news (filter_out (curry (op =) new) olds) |
|
452 |
else |
|
453 |
new :: old :: zip_facts news olds |
|
454 |
||
455 |
fun add_thy_facts_from_thys new old = |
|
456 |
let |
|
457 |
fun add_thy (thy, new_facts) = |
|
458 |
case Symtab.lookup old thy of |
|
459 |
SOME old_facts => Symtab.update (thy, zip_facts old_facts new_facts) |
|
460 |
| NONE => Symtab.update_new (thy, new_facts) |
|
461 |
in old |> Symtab.fold add_thy new end |
|
462 |
||
48308 | 463 |
fun mash_learn_thy ctxt thy timeout = |
48304 | 464 |
let |
465 |
val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt |
|
466 |
val facts = all_non_tautological_facts_of thy css_table |
|
48306 | 467 |
val {thy_facts, ...} = mash_get () |
48304 | 468 |
fun is_old (_, th) = is_fact_in_thy_facts thy_facts (Thm.get_name_hint th) |
469 |
val (old_facts, new_facts) = |
|
470 |
facts |> List.partition is_old ||> sort (thm_ord o pairself snd) |
|
48308 | 471 |
in |
472 |
if null new_facts then |
|
473 |
() |
|
474 |
else |
|
48304 | 475 |
let |
48308 | 476 |
val ths = facts |> map snd |
477 |
val all_names = ths |> map Thm.get_name_hint |
|
48311 | 478 |
fun do_fact ((_, (_, status)), th) (prevs, upds) = |
48308 | 479 |
let |
480 |
val name = Thm.get_name_hint th |
|
481 |
val feats = features_of thy status [prop_of th] |
|
482 |
val deps = isabelle_dependencies_of all_names th |
|
48311 | 483 |
val upd = (name, prevs, feats, deps) |
484 |
in ([name], upd :: upds) end |
|
48308 | 485 |
val parents = parent_facts thy thy_facts |
48311 | 486 |
val (_, upds) = (parents, []) |> fold do_fact new_facts |
48308 | 487 |
val new_thy_facts = new_facts |> thy_facts_from_thms |
488 |
fun trans {dirty_thys, thy_facts} = |
|
48311 | 489 |
(mash_ADD ctxt (rev upds); |
48308 | 490 |
{dirty_thys = dirty_thys, |
491 |
thy_facts = thy_facts |> add_thy_facts_from_thys new_thy_facts}) |
|
492 |
in mash_map trans end |
|
493 |
end |
|
48304 | 494 |
|
48309 | 495 |
fun mash_learn_proof ctxt thy t ths = |
496 |
mash_map (fn state as {dirty_thys, thy_facts} => |
|
497 |
let val deps = ths |> map Thm.get_name_hint in |
|
498 |
if forall (is_fact_in_thy_facts thy_facts) deps then |
|
499 |
let |
|
500 |
val fact = ATP_Util.timestamp () (* should be fairly fresh *) |
|
501 |
val access = accessibility_of thy thy_facts |
|
502 |
val feats = features_of thy General [t] |
|
503 |
in |
|
504 |
mash_ADD ctxt [(fact, access, feats, deps)]; |
|
505 |
{dirty_thys = dirty_thys, thy_facts = thy_facts} |
|
506 |
end |
|
507 |
else |
|
508 |
state |
|
509 |
end) |
|
48249 | 510 |
|
48293 | 511 |
fun relevant_facts ctxt params prover max_facts |
48298 | 512 |
({add, only, ...} : fact_override) hyp_ts concl_t facts = |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
513 |
if only then |
48289 | 514 |
facts |
48293 | 515 |
else if max_facts <= 0 then |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
516 |
[] |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
517 |
else |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
518 |
let |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
519 |
val add_ths = Attrib.eval_thms ctxt add |
48292 | 520 |
fun prepend_facts ths accepts = |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
521 |
((facts |> filter (member Thm.eq_thm_prop ths o snd)) @ |
48292 | 522 |
(accepts |> filter_out (member Thm.eq_thm_prop ths o snd))) |
48293 | 523 |
|> take max_facts |
48298 | 524 |
val iter_facts = |
525 |
iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts |
|
526 |
concl_t facts |
|
48312 | 527 |
val (mash_facts, mash_rejects) = |
48311 | 528 |
facts |> mash_suggest_facts ctxt params prover hyp_ts concl_t |
529 |
|> chop max_facts |
|
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
530 |
in |
48312 | 531 |
mesh_facts max_facts iter_facts mash_facts mash_rejects |
48288
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
532 |
|> not (null add_ths) ? prepend_facts add_ths |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
533 |
end |
255c6e1fd505
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents:
48251
diff
changeset
|
534 |
|
48248 | 535 |
end; |