blanchet@35826
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
|
wenzelm@33309
|
2 |
Author: Jia Meng, Cambridge University Computer Laboratory, NICTA
|
wenzelm@33309
|
3 |
*)
|
paulson@15452
|
4 |
|
blanchet@35826
|
5 |
signature SLEDGEHAMMER_FACT_FILTER =
|
wenzelm@16802
|
6 |
sig
|
blanchet@35865
|
7 |
type classrel_clause = Sledgehammer_FOL_Clause.classrel_clause
|
blanchet@35865
|
8 |
type arity_clause = Sledgehammer_FOL_Clause.arity_clause
|
blanchet@35826
|
9 |
type axiom_name = Sledgehammer_HOL_Clause.axiom_name
|
blanchet@35865
|
10 |
type hol_clause = Sledgehammer_HOL_Clause.hol_clause
|
blanchet@35865
|
11 |
type hol_clause_id = Sledgehammer_HOL_Clause.hol_clause_id
|
blanchet@35966
|
12 |
type relevance_override =
|
blanchet@35966
|
13 |
{add: Facts.ref list,
|
blanchet@35966
|
14 |
del: Facts.ref list,
|
blanchet@35966
|
15 |
only: bool}
|
blanchet@35966
|
16 |
|
paulson@22989
|
17 |
val tvar_classes_of_terms : term list -> string list
|
paulson@22989
|
18 |
val tfree_classes_of_terms : term list -> string list
|
paulson@22989
|
19 |
val type_consts_of_terms : theory -> term list -> string list
|
blanchet@35963
|
20 |
val get_relevant_facts :
|
blanchet@35966
|
21 |
real -> bool option -> bool -> int -> bool -> relevance_override
|
blanchet@35963
|
22 |
-> Proof.context * (thm list * 'a) -> thm list
|
blanchet@35963
|
23 |
-> (thm * (string * int)) list
|
blanchet@35963
|
24 |
val prepare_clauses : bool option -> bool -> thm list -> thm list ->
|
blanchet@35865
|
25 |
(thm * (axiom_name * hol_clause_id)) list ->
|
blanchet@35865
|
26 |
(thm * (axiom_name * hol_clause_id)) list -> theory ->
|
blanchet@35826
|
27 |
axiom_name vector *
|
blanchet@35865
|
28 |
(hol_clause list * hol_clause list * hol_clause list *
|
blanchet@35865
|
29 |
hol_clause list * classrel_clause list * arity_clause list)
|
paulson@15347
|
30 |
end;
|
paulson@15347
|
31 |
|
blanchet@35826
|
32 |
structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
|
paulson@15347
|
33 |
struct
|
paulson@15347
|
34 |
|
blanchet@35865
|
35 |
open Sledgehammer_FOL_Clause
|
blanchet@35865
|
36 |
open Sledgehammer_Fact_Preprocessor
|
blanchet@35865
|
37 |
open Sledgehammer_HOL_Clause
|
blanchet@35826
|
38 |
|
blanchet@35966
|
39 |
type relevance_override =
|
blanchet@35966
|
40 |
{add: Facts.ref list,
|
blanchet@35966
|
41 |
del: Facts.ref list,
|
blanchet@35966
|
42 |
only: bool}
|
paulson@21070
|
43 |
|
mengj@19194
|
44 |
(********************************************************************)
|
mengj@19194
|
45 |
(* some settings for both background automatic ATP calling procedure*)
|
mengj@19194
|
46 |
(* and also explicit ATP invocation methods *)
|
mengj@19194
|
47 |
(********************************************************************)
|
mengj@19194
|
48 |
|
mengj@19194
|
49 |
(*** background linkup ***)
|
wenzelm@28477
|
50 |
val run_blacklist_filter = true;
|
paulson@21224
|
51 |
|
paulson@24287
|
52 |
(*** relevance filter parameters ***)
|
wenzelm@28477
|
53 |
val convergence = 3.2; (*Higher numbers allow longer inference chains*)
|
wenzelm@28477
|
54 |
|
wenzelm@28477
|
55 |
(***************************************************************)
|
wenzelm@28477
|
56 |
(* Relevance Filtering *)
|
wenzelm@28477
|
57 |
(***************************************************************)
|
mengj@19194
|
58 |
|
blanchet@35865
|
59 |
fun strip_Trueprop (@{const Trueprop} $ t) = t
|
paulson@24958
|
60 |
| strip_Trueprop t = t;
|
mengj@19194
|
61 |
|
paulson@24287
|
62 |
(*A surprising number of theorems contain only a few significant constants.
|
paulson@24287
|
63 |
These include all induction rules, and other general theorems. Filtering
|
paulson@24287
|
64 |
theorems in clause form reveals these complexities in the form of Skolem
|
paulson@24287
|
65 |
functions. If we were instead to filter theorems in their natural form,
|
paulson@24287
|
66 |
some other method of measuring theorem complexity would become necessary.*)
|
paulson@24287
|
67 |
|
paulson@24287
|
68 |
fun log_weight2 (x:real) = 1.0 + 2.0/Math.ln (x+1.0);
|
paulson@24287
|
69 |
|
paulson@24287
|
70 |
(*The default seems best in practice. A constant function of one ignores
|
paulson@24287
|
71 |
the constant frequencies.*)
|
wenzelm@28477
|
72 |
val weight_fn = log_weight2;
|
paulson@24287
|
73 |
|
paulson@24287
|
74 |
|
paulson@24287
|
75 |
(*Including equality in this list might be expected to stop rules like subset_antisym from
|
paulson@24287
|
76 |
being chosen, but for some reason filtering works better with them listed. The
|
paulson@24287
|
77 |
logical signs All, Ex, &, and --> are omitted because any remaining occurrrences
|
paulson@24287
|
78 |
must be within comprehensions.*)
|
blanchet@35865
|
79 |
val standard_consts =
|
blanchet@35865
|
80 |
[@{const_name Trueprop}, @{const_name "==>"}, @{const_name all},
|
blanchet@35865
|
81 |
@{const_name "=="}, @{const_name "op |"}, @{const_name Not},
|
blanchet@35865
|
82 |
@{const_name "op ="}];
|
paulson@24287
|
83 |
|
paulson@24287
|
84 |
|
paulson@24287
|
85 |
(*** constants with types ***)
|
paulson@24287
|
86 |
|
paulson@24287
|
87 |
(*An abstraction of Isabelle types*)
|
paulson@24287
|
88 |
datatype const_typ = CTVar | CType of string * const_typ list
|
paulson@24287
|
89 |
|
paulson@24287
|
90 |
(*Is the second type an instance of the first one?*)
|
paulson@24287
|
91 |
fun match_type (CType(con1,args1)) (CType(con2,args2)) =
|
paulson@24287
|
92 |
con1=con2 andalso match_types args1 args2
|
paulson@24287
|
93 |
| match_type CTVar _ = true
|
paulson@24287
|
94 |
| match_type _ CTVar = false
|
paulson@24287
|
95 |
and match_types [] [] = true
|
paulson@24287
|
96 |
| match_types (a1::as1) (a2::as2) = match_type a1 a2 andalso match_types as1 as2;
|
paulson@24287
|
97 |
|
paulson@24287
|
98 |
(*Is there a unifiable constant?*)
|
paulson@24287
|
99 |
fun uni_mem gctab (c,c_typ) =
|
paulson@24287
|
100 |
case Symtab.lookup gctab c of
|
paulson@24287
|
101 |
NONE => false
|
paulson@24287
|
102 |
| SOME ctyps_list => exists (match_types c_typ) ctyps_list;
|
paulson@24287
|
103 |
|
paulson@24287
|
104 |
(*Maps a "real" type to a const_typ*)
|
paulson@24287
|
105 |
fun const_typ_of (Type (c,typs)) = CType (c, map const_typ_of typs)
|
paulson@24287
|
106 |
| const_typ_of (TFree _) = CTVar
|
paulson@24287
|
107 |
| const_typ_of (TVar _) = CTVar
|
paulson@24287
|
108 |
|
paulson@24287
|
109 |
(*Pairs a constant with the list of its type instantiations (using const_typ)*)
|
paulson@24287
|
110 |
fun const_with_typ thy (c,typ) =
|
paulson@24287
|
111 |
let val tvars = Sign.const_typargs thy (c,typ)
|
paulson@24287
|
112 |
in (c, map const_typ_of tvars) end
|
paulson@24287
|
113 |
handle TYPE _ => (c,[]); (*Variable (locale constant): monomorphic*)
|
paulson@24287
|
114 |
|
paulson@24287
|
115 |
(*Add a const/type pair to the table, but a [] entry means a standard connective,
|
paulson@24287
|
116 |
which we ignore.*)
|
paulson@24287
|
117 |
fun add_const_typ_table ((c,ctyps), tab) =
|
paulson@24287
|
118 |
Symtab.map_default (c, [ctyps]) (fn [] => [] | ctyps_list => insert (op =) ctyps ctyps_list)
|
paulson@24287
|
119 |
tab;
|
paulson@24287
|
120 |
|
paulson@24287
|
121 |
(*Free variables are included, as well as constants, to handle locales*)
|
paulson@24287
|
122 |
fun add_term_consts_typs_rm thy (Const(c, typ), tab) =
|
paulson@24287
|
123 |
add_const_typ_table (const_with_typ thy (c,typ), tab)
|
paulson@24287
|
124 |
| add_term_consts_typs_rm thy (Free(c, typ), tab) =
|
paulson@24287
|
125 |
add_const_typ_table (const_with_typ thy (c,typ), tab)
|
paulson@24287
|
126 |
| add_term_consts_typs_rm thy (t $ u, tab) =
|
paulson@24287
|
127 |
add_term_consts_typs_rm thy (t, add_term_consts_typs_rm thy (u, tab))
|
paulson@24287
|
128 |
| add_term_consts_typs_rm thy (Abs(_,_,t), tab) = add_term_consts_typs_rm thy (t, tab)
|
wenzelm@32994
|
129 |
| add_term_consts_typs_rm _ (_, tab) = tab;
|
paulson@24287
|
130 |
|
paulson@24287
|
131 |
(*The empty list here indicates that the constant is being ignored*)
|
paulson@24287
|
132 |
fun add_standard_const (s,tab) = Symtab.update (s,[]) tab;
|
paulson@24287
|
133 |
|
paulson@24287
|
134 |
val null_const_tab : const_typ list list Symtab.table =
|
wenzelm@30190
|
135 |
List.foldl add_standard_const Symtab.empty standard_consts;
|
paulson@24287
|
136 |
|
wenzelm@30190
|
137 |
fun get_goal_consts_typs thy = List.foldl (add_term_consts_typs_rm thy) null_const_tab;
|
paulson@24287
|
138 |
|
paulson@24287
|
139 |
(*Inserts a dummy "constant" referring to the theory name, so that relevance
|
paulson@24287
|
140 |
takes the given theory into account.*)
|
blanchet@35963
|
141 |
fun const_prop_of add_theory_const th =
|
blanchet@35963
|
142 |
if add_theory_const then
|
paulson@24287
|
143 |
let val name = Context.theory_name (theory_of_thm th)
|
paulson@24287
|
144 |
val t = Const (name ^ ". 1", HOLogic.boolT)
|
paulson@24287
|
145 |
in t $ prop_of th end
|
paulson@24287
|
146 |
else prop_of th;
|
paulson@24287
|
147 |
|
paulson@24287
|
148 |
(**** Constant / Type Frequencies ****)
|
paulson@24287
|
149 |
|
paulson@24287
|
150 |
(*A two-dimensional symbol table counts frequencies of constants. It's keyed first by
|
paulson@24287
|
151 |
constant name and second by its list of type instantiations. For the latter, we need
|
paulson@24287
|
152 |
a linear ordering on type const_typ list.*)
|
paulson@24287
|
153 |
|
paulson@24287
|
154 |
local
|
paulson@24287
|
155 |
|
paulson@24287
|
156 |
fun cons_nr CTVar = 0
|
paulson@24287
|
157 |
| cons_nr (CType _) = 1;
|
paulson@24287
|
158 |
|
paulson@24287
|
159 |
in
|
paulson@24287
|
160 |
|
paulson@24287
|
161 |
fun const_typ_ord TU =
|
paulson@24287
|
162 |
case TU of
|
paulson@24287
|
163 |
(CType (a, Ts), CType (b, Us)) =>
|
paulson@24287
|
164 |
(case fast_string_ord(a,b) of EQUAL => dict_ord const_typ_ord (Ts,Us) | ord => ord)
|
paulson@24287
|
165 |
| (T, U) => int_ord (cons_nr T, cons_nr U);
|
paulson@24287
|
166 |
|
paulson@24287
|
167 |
end;
|
paulson@24287
|
168 |
|
wenzelm@31971
|
169 |
structure CTtab = Table(type key = const_typ list val ord = dict_ord const_typ_ord);
|
paulson@24287
|
170 |
|
blanchet@35963
|
171 |
fun count_axiom_consts add_theory_const thy ((thm,_), tab) =
|
paulson@24287
|
172 |
let fun count_const (a, T, tab) =
|
wenzelm@32960
|
173 |
let val (c, cts) = const_with_typ thy (a,T)
|
wenzelm@32960
|
174 |
in (*Two-dimensional table update. Constant maps to types maps to count.*)
|
wenzelm@32960
|
175 |
Symtab.map_default (c, CTtab.empty)
|
wenzelm@32960
|
176 |
(CTtab.map_default (cts,0) (fn n => n+1)) tab
|
wenzelm@32960
|
177 |
end
|
paulson@24287
|
178 |
fun count_term_consts (Const(a,T), tab) = count_const(a,T,tab)
|
wenzelm@32960
|
179 |
| count_term_consts (Free(a,T), tab) = count_const(a,T,tab)
|
wenzelm@32960
|
180 |
| count_term_consts (t $ u, tab) =
|
wenzelm@32960
|
181 |
count_term_consts (t, count_term_consts (u, tab))
|
wenzelm@32960
|
182 |
| count_term_consts (Abs(_,_,t), tab) = count_term_consts (t, tab)
|
wenzelm@32960
|
183 |
| count_term_consts (_, tab) = tab
|
blanchet@35963
|
184 |
in count_term_consts (const_prop_of add_theory_const thm, tab) end;
|
paulson@24287
|
185 |
|
paulson@24287
|
186 |
|
paulson@24287
|
187 |
(**** Actual Filtering Code ****)
|
paulson@24287
|
188 |
|
paulson@24287
|
189 |
(*The frequency of a constant is the sum of those of all instances of its type.*)
|
paulson@24287
|
190 |
fun const_frequency ctab (c, cts) =
|
paulson@24287
|
191 |
let val pairs = CTtab.dest (the (Symtab.lookup ctab c))
|
paulson@24287
|
192 |
fun add ((cts',m), n) = if match_types cts cts' then m+n else n
|
paulson@24287
|
193 |
in List.foldl add 0 pairs end;
|
paulson@24287
|
194 |
|
paulson@24287
|
195 |
(*Add in a constant's weight, as determined by its frequency.*)
|
paulson@24287
|
196 |
fun add_ct_weight ctab ((c,T), w) =
|
wenzelm@28477
|
197 |
w + weight_fn (real (const_frequency ctab (c,T)));
|
paulson@24287
|
198 |
|
paulson@24287
|
199 |
(*Relevant constants are weighted according to frequency,
|
paulson@24287
|
200 |
but irrelevant constants are simply counted. Otherwise, Skolem functions,
|
paulson@24287
|
201 |
which are rare, would harm a clause's chances of being picked.*)
|
paulson@24287
|
202 |
fun clause_weight ctab gctyps consts_typs =
|
paulson@24287
|
203 |
let val rel = filter (uni_mem gctyps) consts_typs
|
paulson@24287
|
204 |
val rel_weight = List.foldl (add_ct_weight ctab) 0.0 rel
|
paulson@24287
|
205 |
in
|
wenzelm@32960
|
206 |
rel_weight / (rel_weight + real (length consts_typs - length rel))
|
paulson@24287
|
207 |
end;
|
paulson@24287
|
208 |
|
paulson@24287
|
209 |
(*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*)
|
wenzelm@30190
|
210 |
fun add_expand_pairs (x,ys) xys = List.foldl (fn (y,acc) => (x,y)::acc) xys ys;
|
paulson@24287
|
211 |
|
paulson@24287
|
212 |
fun consts_typs_of_term thy t =
|
paulson@24287
|
213 |
let val tab = add_term_consts_typs_rm thy (t, null_const_tab)
|
paulson@24287
|
214 |
in Symtab.fold add_expand_pairs tab [] end;
|
paulson@24287
|
215 |
|
blanchet@35963
|
216 |
fun pair_consts_typs_axiom add_theory_const thy (p as (thm, _)) =
|
blanchet@35963
|
217 |
(p, (consts_typs_of_term thy (const_prop_of add_theory_const thm)));
|
paulson@24287
|
218 |
|
paulson@24287
|
219 |
exception ConstFree;
|
paulson@24287
|
220 |
fun dest_ConstFree (Const aT) = aT
|
paulson@24287
|
221 |
| dest_ConstFree (Free aT) = aT
|
paulson@24287
|
222 |
| dest_ConstFree _ = raise ConstFree;
|
paulson@24287
|
223 |
|
paulson@24287
|
224 |
(*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
|
wenzelm@32994
|
225 |
fun defines thy thm gctypes =
|
paulson@24287
|
226 |
let val tm = prop_of thm
|
wenzelm@32960
|
227 |
fun defs lhs rhs =
|
paulson@24287
|
228 |
let val (rator,args) = strip_comb lhs
|
wenzelm@32960
|
229 |
val ct = const_with_typ thy (dest_ConstFree rator)
|
haftmann@33037
|
230 |
in
|
haftmann@33037
|
231 |
forall is_Var args andalso uni_mem gctypes ct andalso
|
haftmann@33038
|
232 |
subset (op =) (Term.add_vars rhs [], Term.add_vars lhs [])
|
paulson@24287
|
233 |
end
|
wenzelm@32960
|
234 |
handle ConstFree => false
|
paulson@24287
|
235 |
in
|
blanchet@35963
|
236 |
case tm of
|
blanchet@35963
|
237 |
@{const Trueprop} $ (Const (@{const_name "op ="}, _) $ lhs $ rhs) =>
|
blanchet@35963
|
238 |
defs lhs rhs
|
blanchet@35963
|
239 |
| _ => false
|
paulson@24287
|
240 |
end;
|
paulson@24287
|
241 |
|
paulson@24287
|
242 |
type annotd_cls = (thm * (string * int)) * ((string * const_typ list) list);
|
paulson@24287
|
243 |
|
paulson@24287
|
244 |
(*For a reverse sort, putting the largest values first.*)
|
paulson@24287
|
245 |
fun compare_pairs ((_,w1),(_,w2)) = Real.compare (w2,w1);
|
paulson@24287
|
246 |
|
paulson@24287
|
247 |
(*Limit the number of new clauses, to prevent runaway acceptance.*)
|
wenzelm@28477
|
248 |
fun take_best max_new (newpairs : (annotd_cls*real) list) =
|
paulson@24287
|
249 |
let val nnew = length newpairs
|
paulson@24287
|
250 |
in
|
wenzelm@28477
|
251 |
if nnew <= max_new then (map #1 newpairs, [])
|
paulson@24287
|
252 |
else
|
paulson@24287
|
253 |
let val cls = sort compare_pairs newpairs
|
wenzelm@28477
|
254 |
val accepted = List.take (cls, max_new)
|
paulson@24287
|
255 |
in
|
blanchet@35865
|
256 |
trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^
|
wenzelm@32960
|
257 |
", exceeds the limit of " ^ Int.toString (max_new)));
|
blanchet@35865
|
258 |
trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
|
blanchet@35865
|
259 |
trace_msg (fn () => "Actually passed: " ^
|
paulson@24287
|
260 |
space_implode ", " (map (fn (((_,(name,_)),_),_) => name) accepted));
|
paulson@24287
|
261 |
|
wenzelm@32960
|
262 |
(map #1 accepted, map #1 (List.drop (cls, max_new)))
|
paulson@24287
|
263 |
end
|
paulson@24287
|
264 |
end;
|
paulson@24287
|
265 |
|
blanchet@35966
|
266 |
fun relevant_clauses ctxt follow_defs max_new
|
blanchet@35966
|
267 |
(relevance_override as {add, del, only}) thy ctab p
|
blanchet@35966
|
268 |
rel_consts =
|
blanchet@35966
|
269 |
let val add_thms = maps (ProofContext.get_fact ctxt) add
|
blanchet@35966
|
270 |
val del_thms = maps (ProofContext.get_fact ctxt) del
|
blanchet@35966
|
271 |
fun relevant ([],_) [] = [] : (thm * (string * int)) list (*Nothing added this iteration*)
|
wenzelm@32960
|
272 |
| relevant (newpairs,rejects) [] =
|
wenzelm@32960
|
273 |
let val (newrels,more_rejects) = take_best max_new newpairs
|
wenzelm@32960
|
274 |
val new_consts = maps #2 newrels
|
wenzelm@32960
|
275 |
val rel_consts' = List.foldl add_const_typ_table rel_consts new_consts
|
wenzelm@32960
|
276 |
val newp = p + (1.0-p) / convergence
|
wenzelm@32960
|
277 |
in
|
blanchet@35865
|
278 |
trace_msg (fn () => "relevant this iteration: " ^ Int.toString (length newrels));
|
blanchet@35966
|
279 |
map #1 newrels @
|
blanchet@35966
|
280 |
relevant_clauses ctxt follow_defs max_new relevance_override thy ctab
|
blanchet@35966
|
281 |
newp rel_consts' (more_rejects @ rejects)
|
wenzelm@32960
|
282 |
end
|
blanchet@35966
|
283 |
| relevant (newrels, rejects)
|
blanchet@35966
|
284 |
((ax as (clsthm as (thm, (name, n)), consts_typs)) :: axs) =
|
blanchet@35966
|
285 |
let
|
blanchet@35966
|
286 |
val weight = if member Thm.eq_thm del_thms thm then 0.0
|
blanchet@35966
|
287 |
else if member Thm.eq_thm add_thms thm then 1.0
|
blanchet@35966
|
288 |
else if only then 0.0
|
blanchet@35966
|
289 |
else clause_weight ctab rel_consts consts_typs
|
wenzelm@32960
|
290 |
in
|
wenzelm@32994
|
291 |
if p <= weight orelse (follow_defs andalso defines thy (#1 clsthm) rel_consts)
|
blanchet@35865
|
292 |
then (trace_msg (fn () => (name ^ " clause " ^ Int.toString n ^
|
wenzelm@32960
|
293 |
" passes: " ^ Real.toString weight));
|
wenzelm@32960
|
294 |
relevant ((ax,weight)::newrels, rejects) axs)
|
wenzelm@32960
|
295 |
else relevant (newrels, ax::rejects) axs
|
wenzelm@32960
|
296 |
end
|
blanchet@35865
|
297 |
in trace_msg (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p));
|
blanchet@35966
|
298 |
relevant ([],[])
|
paulson@24287
|
299 |
end;
|
wenzelm@32960
|
300 |
|
blanchet@35966
|
301 |
fun relevance_filter ctxt relevance_threshold follow_defs max_new
|
blanchet@35966
|
302 |
add_theory_const relevance_override thy axioms goals =
|
blanchet@35966
|
303 |
if relevance_threshold > 0.0 then
|
blanchet@35963
|
304 |
let
|
blanchet@35963
|
305 |
val const_tab = List.foldl (count_axiom_consts add_theory_const thy)
|
blanchet@35963
|
306 |
Symtab.empty axioms
|
paulson@24287
|
307 |
val goal_const_tab = get_goal_consts_typs thy goals
|
blanchet@35963
|
308 |
val _ =
|
blanchet@35963
|
309 |
trace_msg (fn () => "Initial constants: " ^
|
blanchet@35963
|
310 |
commas (Symtab.keys goal_const_tab))
|
blanchet@35963
|
311 |
val relevant =
|
blanchet@35966
|
312 |
relevant_clauses ctxt follow_defs max_new relevance_override thy
|
blanchet@35966
|
313 |
const_tab relevance_threshold goal_const_tab
|
blanchet@35963
|
314 |
(map (pair_consts_typs_axiom add_theory_const thy)
|
blanchet@35963
|
315 |
axioms)
|
blanchet@35963
|
316 |
in
|
blanchet@35963
|
317 |
trace_msg (fn () => "Total relevant: " ^ Int.toString (length relevant));
|
blanchet@35963
|
318 |
relevant
|
blanchet@35963
|
319 |
end
|
blanchet@35963
|
320 |
else
|
blanchet@35963
|
321 |
axioms;
|
paulson@24287
|
322 |
|
paulson@24287
|
323 |
(***************************************************************)
|
mengj@19768
|
324 |
(* Retrieving and filtering lemmas *)
|
mengj@19768
|
325 |
(***************************************************************)
|
mengj@19768
|
326 |
|
paulson@33022
|
327 |
(*** retrieve lemmas and filter them ***)
|
mengj@19768
|
328 |
|
mengj@19768
|
329 |
(*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute*)
|
mengj@19768
|
330 |
|
paulson@22382
|
331 |
fun setinsert (x,s) = Symtab.update (x,()) s;
|
mengj@19768
|
332 |
|
paulson@20757
|
333 |
(*Reject theorems with names like "List.filter.filter_list_def" or
|
paulson@21690
|
334 |
"Accessible_Part.acc.defs", as these are definitions arising from packages.*)
|
paulson@20757
|
335 |
fun is_package_def a =
|
wenzelm@30364
|
336 |
let val names = Long_Name.explode a
|
paulson@21690
|
337 |
in
|
paulson@21690
|
338 |
length names > 2 andalso
|
paulson@21690
|
339 |
not (hd names = "local") andalso
|
paulson@21690
|
340 |
String.isSuffix "_def" a orelse String.isSuffix "_defs" a
|
paulson@21690
|
341 |
end;
|
paulson@20757
|
342 |
|
mengj@19768
|
343 |
(** a hash function from Term.term to int, and also a hash table **)
|
mengj@19768
|
344 |
val xor_words = List.foldl Word.xorb 0w0;
|
mengj@19768
|
345 |
|
mengj@19768
|
346 |
fun hashw_term ((Const(c,_)), w) = Polyhash.hashw_string (c,w)
|
paulson@20661
|
347 |
| hashw_term ((Free(a,_)), w) = Polyhash.hashw_string (a,w)
|
mengj@19768
|
348 |
| hashw_term ((Var(_,_)), w) = w
|
paulson@20661
|
349 |
| hashw_term ((Bound i), w) = Polyhash.hashw_int(i,w)
|
mengj@19768
|
350 |
| hashw_term ((Abs(_,_,t)), w) = hashw_term (t, w)
|
mengj@19768
|
351 |
| hashw_term ((P$Q), w) = hashw_term (Q, (hashw_term (P, w)));
|
mengj@19768
|
352 |
|
blanchet@35865
|
353 |
fun hash_literal (@{const Not} $ P) = Word.notb(hashw_term(P,0w0))
|
paulson@21070
|
354 |
| hash_literal P = hashw_term(P,0w0);
|
mengj@19768
|
355 |
|
paulson@24958
|
356 |
fun hash_term t = Word.toIntX (xor_words (map hash_literal (HOLogic.disjuncts (strip_Trueprop t))));
|
mengj@19768
|
357 |
|
mengj@19768
|
358 |
fun equal_thm (thm1,thm2) = Term.aconv(prop_of thm1, prop_of thm2);
|
paulson@20457
|
359 |
|
paulson@22382
|
360 |
exception HASH_CLAUSE;
|
paulson@22382
|
361 |
|
mengj@19768
|
362 |
(*Create a hash table for clauses, of the given size*)
|
mengj@19768
|
363 |
fun mk_clause_table n =
|
paulson@20457
|
364 |
Polyhash.mkTable (hash_term o prop_of, equal_thm)
|
mengj@19768
|
365 |
(n, HASH_CLAUSE);
|
mengj@19768
|
366 |
|
paulson@20457
|
367 |
(*Use a hash table to eliminate duplicates from xs. Argument is a list of
|
paulson@20868
|
368 |
(thm * (string * int)) tuples. The theorems are hashed into the table. *)
|
wenzelm@21588
|
369 |
fun make_unique xs =
|
paulson@20868
|
370 |
let val ht = mk_clause_table 7000
|
paulson@20457
|
371 |
in
|
blanchet@35865
|
372 |
trace_msg (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses"));
|
wenzelm@21588
|
373 |
app (ignore o Polyhash.peekInsert ht) xs;
|
paulson@20868
|
374 |
Polyhash.listItems ht
|
paulson@20457
|
375 |
end;
|
mengj@19768
|
376 |
|
paulson@20868
|
377 |
(*Remove existing axiom clauses from the conjecture clauses, as this can dramatically
|
paulson@20868
|
378 |
boost an ATP's performance (for some reason)*)
|
wenzelm@21588
|
379 |
fun subtract_cls c_clauses ax_clauses =
|
paulson@20868
|
380 |
let val ht = mk_clause_table 2200
|
wenzelm@33035
|
381 |
fun known x = is_some (Polyhash.peek ht x)
|
paulson@20868
|
382 |
in
|
wenzelm@21588
|
383 |
app (ignore o Polyhash.peekInsert ht) ax_clauses;
|
wenzelm@21588
|
384 |
filter (not o known) c_clauses
|
paulson@20868
|
385 |
end;
|
mengj@19768
|
386 |
|
paulson@21224
|
387 |
fun all_valid_thms ctxt =
|
paulson@22382
|
388 |
let
|
wenzelm@26675
|
389 |
val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt);
|
wenzelm@26278
|
390 |
val local_facts = ProofContext.facts_of ctxt;
|
wenzelm@33641
|
391 |
val full_space =
|
wenzelm@33641
|
392 |
Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts);
|
wenzelm@33641
|
393 |
|
wenzelm@33641
|
394 |
fun valid_facts facts =
|
wenzelm@33641
|
395 |
(facts, []) |-> Facts.fold_static (fn (name, ths0) =>
|
wenzelm@33641
|
396 |
let
|
wenzelm@33641
|
397 |
fun check_thms a =
|
wenzelm@33641
|
398 |
(case try (ProofContext.get_thms ctxt) a of
|
wenzelm@33641
|
399 |
NONE => false
|
wenzelm@33641
|
400 |
| SOME ths1 => Thm.eq_thms (ths0, ths1));
|
wenzelm@33641
|
401 |
|
wenzelm@33641
|
402 |
val name1 = Facts.extern facts name;
|
wenzelm@33641
|
403 |
val name2 = Name_Space.extern full_space name;
|
blanchet@35865
|
404 |
val ths = filter_out bad_for_atp ths0;
|
wenzelm@33641
|
405 |
in
|
wenzelm@33641
|
406 |
if Facts.is_concealed facts name orelse null ths orelse
|
wenzelm@33641
|
407 |
run_blacklist_filter andalso is_package_def name then I
|
wenzelm@33641
|
408 |
else
|
wenzelm@33641
|
409 |
(case find_first check_thms [name1, name2, name] of
|
wenzelm@33641
|
410 |
NONE => I
|
wenzelm@33641
|
411 |
| SOME a => cons (a, ths))
|
wenzelm@33641
|
412 |
end);
|
wenzelm@26675
|
413 |
in valid_facts global_facts @ valid_facts local_facts end;
|
paulson@21224
|
414 |
|
wenzelm@33309
|
415 |
fun multi_name a th (n, pairs) =
|
wenzelm@33309
|
416 |
(n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs);
|
paulson@21224
|
417 |
|
wenzelm@33309
|
418 |
fun add_single_names (a, []) pairs = pairs
|
wenzelm@33309
|
419 |
| add_single_names (a, [th]) pairs = (a, th) :: pairs
|
wenzelm@33309
|
420 |
| add_single_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs));
|
paulson@21431
|
421 |
|
paulson@22382
|
422 |
(*Ignore blacklisted basenames*)
|
wenzelm@33309
|
423 |
fun add_multi_names (a, ths) pairs =
|
blanchet@35865
|
424 |
if (Long_Name.base_name a) mem_string multi_base_blacklist then pairs
|
wenzelm@33309
|
425 |
else add_single_names (a, ths) pairs;
|
paulson@21224
|
426 |
|
paulson@21290
|
427 |
fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a;
|
paulson@21290
|
428 |
|
paulson@24286
|
429 |
(*The single theorems go BEFORE the multiple ones. Blacklist is applied to all.*)
|
wenzelm@21588
|
430 |
fun name_thm_pairs ctxt =
|
wenzelm@33309
|
431 |
let
|
wenzelm@33309
|
432 |
val (mults, singles) = List.partition is_multi (all_valid_thms ctxt)
|
blanchet@35865
|
433 |
val ps = [] |> fold add_multi_names mults
|
blanchet@35865
|
434 |
|> fold add_single_names singles
|
paulson@24286
|
435 |
in
|
blanchet@35865
|
436 |
if run_blacklist_filter then
|
blanchet@35865
|
437 |
let
|
blanchet@35865
|
438 |
val blacklist = No_ATPs.get ctxt
|
blanchet@35865
|
439 |
|> map (`Thm.full_prop_of) |> Termtab.make
|
blanchet@35865
|
440 |
val is_blacklisted = Termtab.defined blacklist o Thm.full_prop_of o snd
|
blanchet@35865
|
441 |
in ps |> filter_out is_blacklisted end
|
blanchet@35865
|
442 |
else
|
blanchet@35865
|
443 |
ps
|
paulson@24286
|
444 |
end;
|
paulson@21224
|
445 |
|
wenzelm@32091
|
446 |
fun check_named ("", th) =
|
wenzelm@32091
|
447 |
(warning ("No name for theorem " ^ Display.string_of_thm_without_context th); false)
|
wenzelm@32994
|
448 |
| check_named _ = true;
|
paulson@19894
|
449 |
|
paulson@33022
|
450 |
fun get_all_lemmas ctxt =
|
paulson@19894
|
451 |
let val included_thms =
|
blanchet@35865
|
452 |
tap (fn ths => trace_msg
|
paulson@33022
|
453 |
(fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
|
paulson@33022
|
454 |
(name_thm_pairs ctxt)
|
paulson@19894
|
455 |
in
|
immler@31410
|
456 |
filter check_named included_thms
|
paulson@19894
|
457 |
end;
|
mengj@19768
|
458 |
|
paulson@21290
|
459 |
(***************************************************************)
|
paulson@21290
|
460 |
(* Type Classes Present in the Axiom or Conjecture Clauses *)
|
paulson@21290
|
461 |
(***************************************************************)
|
paulson@21290
|
462 |
|
wenzelm@32952
|
463 |
fun add_classes (sorts, cset) = List.foldl setinsert cset (flat sorts);
|
paulson@21290
|
464 |
|
paulson@21290
|
465 |
(*Remove this trivial type class*)
|
blanchet@35865
|
466 |
fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset;
|
paulson@21290
|
467 |
|
paulson@21290
|
468 |
fun tvar_classes_of_terms ts =
|
wenzelm@29270
|
469 |
let val sorts_list = map (map #2 o OldTerm.term_tvars) ts
|
wenzelm@30190
|
470 |
in Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list)) end;
|
paulson@21290
|
471 |
|
paulson@21290
|
472 |
fun tfree_classes_of_terms ts =
|
wenzelm@29270
|
473 |
let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts
|
wenzelm@30190
|
474 |
in Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list)) end;
|
paulson@20526
|
475 |
|
paulson@21373
|
476 |
(*fold type constructors*)
|
paulson@21373
|
477 |
fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x))
|
wenzelm@32994
|
478 |
| fold_type_consts _ _ x = x;
|
paulson@21373
|
479 |
|
paulson@21373
|
480 |
val add_type_consts_in_type = fold_type_consts setinsert;
|
paulson@21373
|
481 |
|
paulson@21397
|
482 |
(*Type constructors used to instantiate overloaded constants are the only ones needed.*)
|
paulson@21397
|
483 |
fun add_type_consts_in_term thy =
|
paulson@21397
|
484 |
let val const_typargs = Sign.const_typargs thy
|
paulson@21397
|
485 |
fun add_tcs (Const cT) x = fold add_type_consts_in_type (const_typargs cT) x
|
wenzelm@32994
|
486 |
| add_tcs (Abs (_, _, u)) x = add_tcs u x
|
paulson@21397
|
487 |
| add_tcs (t $ u) x = add_tcs t (add_tcs u x)
|
paulson@21397
|
488 |
| add_tcs _ x = x
|
paulson@21397
|
489 |
in add_tcs end
|
paulson@21373
|
490 |
|
paulson@21397
|
491 |
fun type_consts_of_terms thy ts =
|
paulson@21397
|
492 |
Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty);
|
paulson@21373
|
493 |
|
paulson@21373
|
494 |
|
mengj@19194
|
495 |
(***************************************************************)
|
mengj@19194
|
496 |
(* ATP invocation methods setup *)
|
mengj@19194
|
497 |
(***************************************************************)
|
mengj@19194
|
498 |
|
paulson@20526
|
499 |
(*Ensures that no higher-order theorems "leak out"*)
|
paulson@24958
|
500 |
fun restrict_to_logic thy true cls = filter (Meson.is_fol_term thy o prop_of o fst) cls
|
paulson@24958
|
501 |
| restrict_to_logic thy false cls = cls;
|
paulson@20526
|
502 |
|
paulson@21470
|
503 |
(**** Predicates to detect unwanted clauses (prolific or likely to cause unsoundness) ****)
|
paulson@21470
|
504 |
|
paulson@21470
|
505 |
(** Too general means, positive equality literal with a variable X as one operand,
|
paulson@21470
|
506 |
when X does not occur properly in the other operand. This rules out clearly
|
paulson@21470
|
507 |
inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **)
|
wenzelm@21588
|
508 |
|
paulson@21470
|
509 |
fun occurs ix =
|
paulson@21470
|
510 |
let fun occ(Var (jx,_)) = (ix=jx)
|
paulson@21470
|
511 |
| occ(t1$t2) = occ t1 orelse occ t2
|
paulson@21470
|
512 |
| occ(Abs(_,_,t)) = occ t
|
paulson@21470
|
513 |
| occ _ = false
|
paulson@21470
|
514 |
in occ end;
|
paulson@21470
|
515 |
|
haftmann@31723
|
516 |
fun is_recordtype T = not (null (Record.dest_recTs T));
|
paulson@21470
|
517 |
|
paulson@21470
|
518 |
(*Unwanted equalities include
|
paulson@21470
|
519 |
(1) those between a variable that does not properly occur in the second operand,
|
paulson@21470
|
520 |
(2) those between a variable and a record, since these seem to be prolific "cases" thms
|
wenzelm@21588
|
521 |
*)
|
paulson@21470
|
522 |
fun too_general_eqterms (Var (ix,T), t) = not (occurs ix t) orelse is_recordtype T
|
paulson@21470
|
523 |
| too_general_eqterms _ = false;
|
paulson@21470
|
524 |
|
blanchet@35865
|
525 |
fun too_general_equality (Const (@{const_name "op ="}, _) $ x $ y) =
|
paulson@21470
|
526 |
too_general_eqterms (x,y) orelse too_general_eqterms(y,x)
|
paulson@21470
|
527 |
| too_general_equality _ = false;
|
paulson@21470
|
528 |
|
wenzelm@29267
|
529 |
fun has_typed_var tycons = exists_subterm
|
wenzelm@29267
|
530 |
(fn Var (_, Type (a, _)) => member (op =) tycons a | _ => false);
|
paulson@21431
|
531 |
|
paulson@22217
|
532 |
(*Clauses are forbidden to contain variables of these types. The typical reason is that
|
paulson@22217
|
533 |
they lead to unsoundness. Note that "unit" satisfies numerous equations like ?X=().
|
paulson@22217
|
534 |
The resulting clause will have no type constraint, yielding false proofs. Even "bool"
|
paulson@22217
|
535 |
leads to many unsound proofs, though (obviously) only for higher-order problems.*)
|
blanchet@35865
|
536 |
val unwanted_types = [@{type_name unit}, @{type_name bool}];
|
paulson@22217
|
537 |
|
paulson@21470
|
538 |
fun unwanted t =
|
blanchet@35865
|
539 |
t = @{prop True} orelse has_typed_var unwanted_types t orelse
|
paulson@24958
|
540 |
forall too_general_equality (HOLogic.disjuncts (strip_Trueprop t));
|
paulson@21470
|
541 |
|
paulson@21431
|
542 |
(*Clauses containing variables of type "unit" or "bool" are unlikely to be useful and
|
paulson@21431
|
543 |
likely to lead to unsound proofs.*)
|
paulson@22217
|
544 |
fun remove_unwanted_clauses cls = filter (not o unwanted o prop_of o fst) cls;
|
paulson@21431
|
545 |
|
blanchet@35963
|
546 |
fun is_first_order thy higher_order goal_cls =
|
blanchet@35963
|
547 |
case higher_order of
|
blanchet@35963
|
548 |
NONE => forall (Meson.is_fol_term thy) (map prop_of goal_cls)
|
blanchet@35963
|
549 |
| SOME b => not b
|
immler@30536
|
550 |
|
blanchet@35963
|
551 |
fun get_relevant_facts relevance_threshold higher_order follow_defs max_new
|
blanchet@35966
|
552 |
add_theory_const relevance_override
|
blanchet@35966
|
553 |
(ctxt, (chain_ths, th)) goal_cls =
|
immler@30536
|
554 |
let
|
immler@30536
|
555 |
val thy = ProofContext.theory_of ctxt
|
blanchet@35963
|
556 |
val is_FO = is_first_order thy higher_order goal_cls
|
wenzelm@33306
|
557 |
val included_cls = get_all_lemmas ctxt
|
blanchet@35865
|
558 |
|> cnf_rules_pairs thy |> make_unique
|
blanchet@35865
|
559 |
|> restrict_to_logic thy is_FO
|
wenzelm@33306
|
560 |
|> remove_unwanted_clauses
|
mengj@19194
|
561 |
in
|
blanchet@35966
|
562 |
relevance_filter ctxt relevance_threshold follow_defs max_new
|
blanchet@35966
|
563 |
add_theory_const relevance_override thy included_cls
|
blanchet@35966
|
564 |
(map prop_of goal_cls)
|
immler@30536
|
565 |
end;
|
immler@30536
|
566 |
|
immler@31752
|
567 |
(* prepare for passing to writer,
|
immler@31752
|
568 |
create additional clauses based on the information from extra_cls *)
|
blanchet@35963
|
569 |
fun prepare_clauses higher_order dfg goal_cls chain_ths axcls extra_cls thy =
|
immler@31409
|
570 |
let
|
blanchet@35966
|
571 |
(* ### *)
|
blanchet@35966
|
572 |
val _ = priority ("prepare clauses: axiom " ^ Int.toString (length axcls) ^
|
blanchet@35966
|
573 |
", extra " ^ Int.toString (length extra_cls))
|
Philipp@32866
|
574 |
(* add chain thms *)
|
wenzelm@33306
|
575 |
val chain_cls =
|
blanchet@35865
|
576 |
cnf_rules_pairs thy (filter check_named (map pairname chain_ths))
|
Philipp@32866
|
577 |
val axcls = chain_cls @ axcls
|
Philipp@32866
|
578 |
val extra_cls = chain_cls @ extra_cls
|
blanchet@35963
|
579 |
val is_FO = is_first_order thy higher_order goal_cls
|
immler@31752
|
580 |
val ccls = subtract_cls goal_cls extra_cls
|
blanchet@35865
|
581 |
val _ = app (fn th => trace_msg (fn _ => Display.string_of_thm_global thy th)) ccls
|
immler@30536
|
582 |
val ccltms = map prop_of ccls
|
immler@31752
|
583 |
and axtms = map (prop_of o #1) extra_cls
|
immler@30536
|
584 |
val subs = tfree_classes_of_terms ccltms
|
immler@30536
|
585 |
and supers = tvar_classes_of_terms axtms
|
blanchet@35865
|
586 |
and tycons = type_consts_of_terms thy (ccltms @ axtms)
|
immler@30536
|
587 |
(*TFrees in conjecture clauses; TVars in axiom clauses*)
|
blanchet@35865
|
588 |
val conjectures = make_conjecture_clauses dfg thy ccls
|
blanchet@35865
|
589 |
val (_, extra_clauses) = ListPair.unzip (make_axiom_clauses dfg thy extra_cls)
|
blanchet@35865
|
590 |
val (clnames, axiom_clauses) = ListPair.unzip (make_axiom_clauses dfg thy axcls)
|
blanchet@35865
|
591 |
val helper_clauses = get_helper_clauses dfg thy is_FO (conjectures, extra_cls, [])
|
blanchet@35865
|
592 |
val (supers', arity_clauses) = make_arity_clauses_dfg dfg thy tycons supers
|
blanchet@35865
|
593 |
val classrel_clauses = make_classrel_clauses thy subs supers'
|
immler@30536
|
594 |
in
|
immler@31752
|
595 |
(Vector.fromList clnames,
|
immler@31865
|
596 |
(conjectures, axiom_clauses, extra_clauses, helper_clauses, classrel_clauses, arity_clauses))
|
immler@31409
|
597 |
end
|
quigley@15644
|
598 |
|
paulson@15347
|
599 |
end;
|