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