author | immler@in.tum.de |
Thu, 26 Feb 2009 10:13:43 +0100 | |
changeset 30151 | 629f3a92863e |
parent 29270 | 0eade173f77e |
child 30190 | 479806475f3c |
permissions | -rw-r--r-- |
29267 | 1 |
(* Author: Jia Meng, Cambridge University Computer Laboratory, NICTA *) |
15452 | 2 |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
3 |
signature RES_ATP = |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
4 |
sig |
19194 | 5 |
datatype mode = Auto | Fol | Hol |
22989 | 6 |
val tvar_classes_of_terms : term list -> string list |
7 |
val tfree_classes_of_terms : term list -> string list |
|
8 |
val type_consts_of_terms : theory -> term list -> string list |
|
30151 | 9 |
val write_problem_files : bool -> int -> bool |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
10 |
-> (int -> Path.T) -> Proof.context * thm list * thm |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
11 |
-> string list * ResHolClause.axiom_name Vector.vector list |
15347 | 12 |
end; |
13 |
||
22865 | 14 |
structure ResAtp: RES_ATP = |
15347 | 15 |
struct |
16 |
||
21070 | 17 |
|
19194 | 18 |
(********************************************************************) |
19 |
(* some settings for both background automatic ATP calling procedure*) |
|
20 |
(* and also explicit ATP invocation methods *) |
|
21 |
(********************************************************************) |
|
22 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
23 |
(*Translation mode can be auto-detected, or forced to be first-order or higher-order*) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
24 |
datatype mode = Auto | Fol | Hol; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
25 |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
26 |
val linkup_logic_mode = Auto; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
27 |
|
19194 | 28 |
(*** background linkup ***) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
29 |
val run_blacklist_filter = true; |
21224 | 30 |
|
24287 | 31 |
(*** relevance filter parameters ***) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
32 |
val run_relevance_filter = true; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
33 |
val pass_mark = 0.5; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
34 |
val convergence = 3.2; (*Higher numbers allow longer inference chains*) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
35 |
val follow_defs = false; (*Follow definitions. Makes problems bigger.*) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
36 |
val include_all = true; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
37 |
val include_simpset = false; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
38 |
val include_claset = false; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
39 |
val include_atpset = true; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
40 |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
41 |
(***************************************************************) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
42 |
(* Relevance Filtering *) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
43 |
(***************************************************************) |
19194 | 44 |
|
24958 | 45 |
fun strip_Trueprop (Const("Trueprop",_) $ t) = t |
46 |
| strip_Trueprop t = t; |
|
19194 | 47 |
|
24287 | 48 |
(*A surprising number of theorems contain only a few significant constants. |
49 |
These include all induction rules, and other general theorems. Filtering |
|
50 |
theorems in clause form reveals these complexities in the form of Skolem |
|
51 |
functions. If we were instead to filter theorems in their natural form, |
|
52 |
some other method of measuring theorem complexity would become necessary.*) |
|
53 |
||
54 |
fun log_weight2 (x:real) = 1.0 + 2.0/Math.ln (x+1.0); |
|
55 |
||
56 |
(*The default seems best in practice. A constant function of one ignores |
|
57 |
the constant frequencies.*) |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
58 |
val weight_fn = log_weight2; |
24287 | 59 |
|
60 |
||
61 |
(*Including equality in this list might be expected to stop rules like subset_antisym from |
|
62 |
being chosen, but for some reason filtering works better with them listed. The |
|
63 |
logical signs All, Ex, &, and --> are omitted because any remaining occurrrences |
|
64 |
must be within comprehensions.*) |
|
65 |
val standard_consts = ["Trueprop","==>","all","==","op |","Not","op ="]; |
|
66 |
||
67 |
||
68 |
(*** constants with types ***) |
|
69 |
||
70 |
(*An abstraction of Isabelle types*) |
|
71 |
datatype const_typ = CTVar | CType of string * const_typ list |
|
72 |
||
73 |
(*Is the second type an instance of the first one?*) |
|
74 |
fun match_type (CType(con1,args1)) (CType(con2,args2)) = |
|
75 |
con1=con2 andalso match_types args1 args2 |
|
76 |
| match_type CTVar _ = true |
|
77 |
| match_type _ CTVar = false |
|
78 |
and match_types [] [] = true |
|
79 |
| match_types (a1::as1) (a2::as2) = match_type a1 a2 andalso match_types as1 as2; |
|
80 |
||
81 |
(*Is there a unifiable constant?*) |
|
82 |
fun uni_mem gctab (c,c_typ) = |
|
83 |
case Symtab.lookup gctab c of |
|
84 |
NONE => false |
|
85 |
| SOME ctyps_list => exists (match_types c_typ) ctyps_list; |
|
86 |
||
87 |
(*Maps a "real" type to a const_typ*) |
|
88 |
fun const_typ_of (Type (c,typs)) = CType (c, map const_typ_of typs) |
|
89 |
| const_typ_of (TFree _) = CTVar |
|
90 |
| const_typ_of (TVar _) = CTVar |
|
91 |
||
92 |
(*Pairs a constant with the list of its type instantiations (using const_typ)*) |
|
93 |
fun const_with_typ thy (c,typ) = |
|
94 |
let val tvars = Sign.const_typargs thy (c,typ) |
|
95 |
in (c, map const_typ_of tvars) end |
|
96 |
handle TYPE _ => (c,[]); (*Variable (locale constant): monomorphic*) |
|
97 |
||
98 |
(*Add a const/type pair to the table, but a [] entry means a standard connective, |
|
99 |
which we ignore.*) |
|
100 |
fun add_const_typ_table ((c,ctyps), tab) = |
|
101 |
Symtab.map_default (c, [ctyps]) (fn [] => [] | ctyps_list => insert (op =) ctyps ctyps_list) |
|
102 |
tab; |
|
103 |
||
104 |
(*Free variables are included, as well as constants, to handle locales*) |
|
105 |
fun add_term_consts_typs_rm thy (Const(c, typ), tab) = |
|
106 |
add_const_typ_table (const_with_typ thy (c,typ), tab) |
|
107 |
| add_term_consts_typs_rm thy (Free(c, typ), tab) = |
|
108 |
add_const_typ_table (const_with_typ thy (c,typ), tab) |
|
109 |
| add_term_consts_typs_rm thy (t $ u, tab) = |
|
110 |
add_term_consts_typs_rm thy (t, add_term_consts_typs_rm thy (u, tab)) |
|
111 |
| add_term_consts_typs_rm thy (Abs(_,_,t), tab) = add_term_consts_typs_rm thy (t, tab) |
|
112 |
| add_term_consts_typs_rm thy (_, tab) = tab; |
|
113 |
||
114 |
(*The empty list here indicates that the constant is being ignored*) |
|
115 |
fun add_standard_const (s,tab) = Symtab.update (s,[]) tab; |
|
116 |
||
117 |
val null_const_tab : const_typ list list Symtab.table = |
|
118 |
foldl add_standard_const Symtab.empty standard_consts; |
|
119 |
||
120 |
fun get_goal_consts_typs thy = foldl (add_term_consts_typs_rm thy) null_const_tab; |
|
121 |
||
122 |
(*Inserts a dummy "constant" referring to the theory name, so that relevance |
|
123 |
takes the given theory into account.*) |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
124 |
fun const_prop_of theory_const th = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
125 |
if theory_const then |
24287 | 126 |
let val name = Context.theory_name (theory_of_thm th) |
127 |
val t = Const (name ^ ". 1", HOLogic.boolT) |
|
128 |
in t $ prop_of th end |
|
129 |
else prop_of th; |
|
130 |
||
131 |
(**** Constant / Type Frequencies ****) |
|
132 |
||
133 |
(*A two-dimensional symbol table counts frequencies of constants. It's keyed first by |
|
134 |
constant name and second by its list of type instantiations. For the latter, we need |
|
135 |
a linear ordering on type const_typ list.*) |
|
136 |
||
137 |
local |
|
138 |
||
139 |
fun cons_nr CTVar = 0 |
|
140 |
| cons_nr (CType _) = 1; |
|
141 |
||
142 |
in |
|
143 |
||
144 |
fun const_typ_ord TU = |
|
145 |
case TU of |
|
146 |
(CType (a, Ts), CType (b, Us)) => |
|
147 |
(case fast_string_ord(a,b) of EQUAL => dict_ord const_typ_ord (Ts,Us) | ord => ord) |
|
148 |
| (T, U) => int_ord (cons_nr T, cons_nr U); |
|
149 |
||
150 |
end; |
|
151 |
||
152 |
structure CTtab = TableFun(type key = const_typ list val ord = dict_ord const_typ_ord); |
|
153 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
154 |
fun count_axiom_consts theory_const thy ((thm,_), tab) = |
24287 | 155 |
let fun count_const (a, T, tab) = |
156 |
let val (c, cts) = const_with_typ thy (a,T) |
|
157 |
in (*Two-dimensional table update. Constant maps to types maps to count.*) |
|
158 |
Symtab.map_default (c, CTtab.empty) |
|
159 |
(CTtab.map_default (cts,0) (fn n => n+1)) tab |
|
160 |
end |
|
161 |
fun count_term_consts (Const(a,T), tab) = count_const(a,T,tab) |
|
162 |
| count_term_consts (Free(a,T), tab) = count_const(a,T,tab) |
|
163 |
| count_term_consts (t $ u, tab) = |
|
164 |
count_term_consts (t, count_term_consts (u, tab)) |
|
165 |
| count_term_consts (Abs(_,_,t), tab) = count_term_consts (t, tab) |
|
166 |
| count_term_consts (_, tab) = tab |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
167 |
in count_term_consts (const_prop_of theory_const thm, tab) end; |
24287 | 168 |
|
169 |
||
170 |
(**** Actual Filtering Code ****) |
|
171 |
||
172 |
(*The frequency of a constant is the sum of those of all instances of its type.*) |
|
173 |
fun const_frequency ctab (c, cts) = |
|
174 |
let val pairs = CTtab.dest (the (Symtab.lookup ctab c)) |
|
175 |
fun add ((cts',m), n) = if match_types cts cts' then m+n else n |
|
176 |
in List.foldl add 0 pairs end; |
|
177 |
||
178 |
(*Add in a constant's weight, as determined by its frequency.*) |
|
179 |
fun add_ct_weight ctab ((c,T), w) = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
180 |
w + weight_fn (real (const_frequency ctab (c,T))); |
24287 | 181 |
|
182 |
(*Relevant constants are weighted according to frequency, |
|
183 |
but irrelevant constants are simply counted. Otherwise, Skolem functions, |
|
184 |
which are rare, would harm a clause's chances of being picked.*) |
|
185 |
fun clause_weight ctab gctyps consts_typs = |
|
186 |
let val rel = filter (uni_mem gctyps) consts_typs |
|
187 |
val rel_weight = List.foldl (add_ct_weight ctab) 0.0 rel |
|
188 |
in |
|
189 |
rel_weight / (rel_weight + real (length consts_typs - length rel)) |
|
190 |
end; |
|
191 |
||
192 |
(*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*) |
|
193 |
fun add_expand_pairs (x,ys) xys = foldl (fn (y,acc) => (x,y)::acc) xys ys; |
|
194 |
||
195 |
fun consts_typs_of_term thy t = |
|
196 |
let val tab = add_term_consts_typs_rm thy (t, null_const_tab) |
|
197 |
in Symtab.fold add_expand_pairs tab [] end; |
|
198 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
199 |
fun pair_consts_typs_axiom theory_const thy (thm,name) = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
200 |
((thm,name), (consts_typs_of_term thy (const_prop_of theory_const thm))); |
24287 | 201 |
|
202 |
exception ConstFree; |
|
203 |
fun dest_ConstFree (Const aT) = aT |
|
204 |
| dest_ConstFree (Free aT) = aT |
|
205 |
| dest_ConstFree _ = raise ConstFree; |
|
206 |
||
207 |
(*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*) |
|
208 |
fun defines thy (thm,(name,n)) gctypes = |
|
209 |
let val tm = prop_of thm |
|
210 |
fun defs lhs rhs = |
|
211 |
let val (rator,args) = strip_comb lhs |
|
212 |
val ct = const_with_typ thy (dest_ConstFree rator) |
|
213 |
in forall is_Var args andalso uni_mem gctypes ct andalso |
|
214 |
Term.add_vars rhs [] subset Term.add_vars lhs [] |
|
215 |
end |
|
216 |
handle ConstFree => false |
|
217 |
in |
|
218 |
case tm of Const ("Trueprop",_) $ (Const("op =",_) $ lhs $ rhs) => |
|
219 |
defs lhs rhs andalso |
|
220 |
(Output.debug (fn () => "Definition found: " ^ name ^ "_" ^ Int.toString n); true) |
|
221 |
| _ => false |
|
222 |
end; |
|
223 |
||
224 |
type annotd_cls = (thm * (string * int)) * ((string * const_typ list) list); |
|
225 |
||
226 |
(*For a reverse sort, putting the largest values first.*) |
|
227 |
fun compare_pairs ((_,w1),(_,w2)) = Real.compare (w2,w1); |
|
228 |
||
229 |
(*Limit the number of new clauses, to prevent runaway acceptance.*) |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
230 |
fun take_best max_new (newpairs : (annotd_cls*real) list) = |
24287 | 231 |
let val nnew = length newpairs |
232 |
in |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
233 |
if nnew <= max_new then (map #1 newpairs, []) |
24287 | 234 |
else |
235 |
let val cls = sort compare_pairs newpairs |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
236 |
val accepted = List.take (cls, max_new) |
24287 | 237 |
in |
238 |
Output.debug (fn () => ("Number of candidates, " ^ Int.toString nnew ^ |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
239 |
", exceeds the limit of " ^ Int.toString (max_new))); |
24287 | 240 |
Output.debug (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted)))); |
241 |
Output.debug (fn () => "Actually passed: " ^ |
|
242 |
space_implode ", " (map (fn (((_,(name,_)),_),_) => name) accepted)); |
|
243 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
244 |
(map #1 accepted, map #1 (List.drop (cls, max_new))) |
24287 | 245 |
end |
246 |
end; |
|
247 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
248 |
fun relevant_clauses max_new thy ctab p rel_consts = |
24287 | 249 |
let fun relevant ([],_) [] = [] : (thm * (string * int)) list (*Nothing added this iteration*) |
250 |
| relevant (newpairs,rejects) [] = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
251 |
let val (newrels,more_rejects) = take_best max_new newpairs |
24287 | 252 |
val new_consts = List.concat (map #2 newrels) |
253 |
val rel_consts' = foldl add_const_typ_table rel_consts new_consts |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
254 |
val newp = p + (1.0-p) / convergence |
24287 | 255 |
in |
256 |
Output.debug (fn () => ("relevant this iteration: " ^ Int.toString (length newrels))); |
|
257 |
(map #1 newrels) @ |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
258 |
(relevant_clauses max_new thy ctab newp rel_consts' (more_rejects@rejects)) |
24287 | 259 |
end |
260 |
| relevant (newrels,rejects) ((ax as (clsthm as (_,(name,n)),consts_typs)) :: axs) = |
|
261 |
let val weight = clause_weight ctab rel_consts consts_typs |
|
262 |
in |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
263 |
if p <= weight orelse (follow_defs andalso defines thy clsthm rel_consts) |
24287 | 264 |
then (Output.debug (fn () => (name ^ " clause " ^ Int.toString n ^ |
265 |
" passes: " ^ Real.toString weight)); |
|
266 |
relevant ((ax,weight)::newrels, rejects) axs) |
|
267 |
else relevant (newrels, ax::rejects) axs |
|
268 |
end |
|
269 |
in Output.debug (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p)); |
|
270 |
relevant ([],[]) |
|
271 |
end; |
|
272 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
273 |
fun relevance_filter max_new theory_const thy axioms goals = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
274 |
if run_relevance_filter andalso pass_mark >= 0.1 |
24287 | 275 |
then |
276 |
let val _ = Output.debug (fn () => "Start of relevance filtering"); |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
277 |
val const_tab = List.foldl (count_axiom_consts theory_const thy) Symtab.empty axioms |
24287 | 278 |
val goal_const_tab = get_goal_consts_typs thy goals |
279 |
val _ = Output.debug (fn () => ("Initial constants: " ^ |
|
280 |
space_implode ", " (Symtab.keys goal_const_tab))); |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
281 |
val rels = relevant_clauses max_new thy const_tab (pass_mark) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
282 |
goal_const_tab (map (pair_consts_typs_axiom theory_const thy) axioms) |
24287 | 283 |
in |
284 |
Output.debug (fn () => ("Total relevant: " ^ Int.toString (length rels))); |
|
285 |
rels |
|
286 |
end |
|
287 |
else axioms; |
|
288 |
||
289 |
(***************************************************************) |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
290 |
(* Retrieving and filtering lemmas *) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
291 |
(***************************************************************) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
292 |
|
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
293 |
(*** white list and black list of lemmas ***) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
294 |
|
24215 | 295 |
(*The rule subsetI is frequently omitted by the relevance filter. This could be theory data |
296 |
or identified with ATPset (which however is too big currently)*) |
|
297 |
val whitelist = [subsetI]; |
|
21588 | 298 |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
299 |
(*** retrieve lemmas from clasimpset and atpset, may filter them ***) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
300 |
|
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
301 |
(*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute*) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
302 |
|
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
303 |
fun setinsert (x,s) = Symtab.update (x,()) s; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
304 |
|
20757
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
paulson
parents:
20661
diff
changeset
|
305 |
(*Reject theorems with names like "List.filter.filter_list_def" or |
21690
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
306 |
"Accessible_Part.acc.defs", as these are definitions arising from packages.*) |
20757
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
paulson
parents:
20661
diff
changeset
|
307 |
fun is_package_def a = |
21858
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
wenzelm
parents:
21790
diff
changeset
|
308 |
let val names = NameSpace.explode a |
21690
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
309 |
in |
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
310 |
length names > 2 andalso |
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
311 |
not (hd names = "local") andalso |
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
312 |
String.isSuffix "_def" a orelse String.isSuffix "_defs" a |
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents:
21588
diff
changeset
|
313 |
end; |
20757
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
paulson
parents:
20661
diff
changeset
|
314 |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
315 |
(** a hash function from Term.term to int, and also a hash table **) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
316 |
val xor_words = List.foldl Word.xorb 0w0; |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
317 |
|
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
318 |
fun hashw_term ((Const(c,_)), w) = Polyhash.hashw_string (c,w) |
20661
46832fee1215
Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
paulson
parents:
20643
diff
changeset
|
319 |
| hashw_term ((Free(a,_)), w) = Polyhash.hashw_string (a,w) |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
320 |
| hashw_term ((Var(_,_)), w) = w |
20661
46832fee1215
Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
paulson
parents:
20643
diff
changeset
|
321 |
| hashw_term ((Bound i), w) = Polyhash.hashw_int(i,w) |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
322 |
| hashw_term ((Abs(_,_,t)), w) = hashw_term (t, w) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
323 |
| hashw_term ((P$Q), w) = hashw_term (Q, (hashw_term (P, w))); |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
324 |
|
21070 | 325 |
fun hash_literal (Const("Not",_)$P) = Word.notb(hashw_term(P,0w0)) |
326 |
| hash_literal P = hashw_term(P,0w0); |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
327 |
|
24958 | 328 |
fun hash_term t = Word.toIntX (xor_words (map hash_literal (HOLogic.disjuncts (strip_Trueprop t)))); |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
329 |
|
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
330 |
fun equal_thm (thm1,thm2) = Term.aconv(prop_of thm1, prop_of thm2); |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
331 |
|
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
332 |
exception HASH_CLAUSE; |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
333 |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
334 |
(*Create a hash table for clauses, of the given size*) |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
335 |
fun mk_clause_table n = |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
336 |
Polyhash.mkTable (hash_term o prop_of, equal_thm) |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
337 |
(n, HASH_CLAUSE); |
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
338 |
|
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
339 |
(*Use a hash table to eliminate duplicates from xs. Argument is a list of |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
340 |
(thm * (string * int)) tuples. The theorems are hashed into the table. *) |
21588 | 341 |
fun make_unique xs = |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
342 |
let val ht = mk_clause_table 7000 |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
343 |
in |
22130 | 344 |
Output.debug (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses")); |
21588 | 345 |
app (ignore o Polyhash.peekInsert ht) xs; |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
346 |
Polyhash.listItems ht |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
347 |
end; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
348 |
|
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
349 |
(*Remove existing axiom clauses from the conjecture clauses, as this can dramatically |
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
350 |
boost an ATP's performance (for some reason)*) |
21588 | 351 |
fun subtract_cls c_clauses ax_clauses = |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
352 |
let val ht = mk_clause_table 2200 |
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
353 |
fun known x = isSome (Polyhash.peek ht x) |
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
354 |
in |
21588 | 355 |
app (ignore o Polyhash.peekInsert ht) ax_clauses; |
356 |
filter (not o known) c_clauses |
|
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
357 |
end; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
358 |
|
26675 | 359 |
fun valid_facts facts = |
26691
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents:
26675
diff
changeset
|
360 |
Facts.fold_static (fn (name, ths) => |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
361 |
if run_blacklist_filter andalso is_package_def name then I |
26691
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents:
26675
diff
changeset
|
362 |
else |
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents:
26675
diff
changeset
|
363 |
let val xname = Facts.extern facts name in |
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents:
26675
diff
changeset
|
364 |
if NameSpace.is_hidden xname then I |
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents:
26675
diff
changeset
|
365 |
else cons (xname, filter_out ResAxioms.bad_for_atp ths) |
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
wenzelm
parents:
26675
diff
changeset
|
366 |
end) facts []; |
26675 | 367 |
|
21224 | 368 |
fun all_valid_thms ctxt = |
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
369 |
let |
26675 | 370 |
val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt); |
26278 | 371 |
val local_facts = ProofContext.facts_of ctxt; |
26675 | 372 |
in valid_facts global_facts @ valid_facts local_facts end; |
21224 | 373 |
|
21588 | 374 |
fun multi_name a (th, (n,pairs)) = |
21224 | 375 |
(n+1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs) |
376 |
||
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
377 |
fun add_single_names ((a, []), pairs) = pairs |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
378 |
| add_single_names ((a, [th]), pairs) = (a,th)::pairs |
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
379 |
| add_single_names ((a, ths), pairs) = #2 (foldl (multi_name a) (1,pairs) ths); |
21431
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
380 |
|
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
381 |
(*Ignore blacklisted basenames*) |
21588 | 382 |
fun add_multi_names ((a, ths), pairs) = |
24854 | 383 |
if (Sign.base_name a) mem_string ResAxioms.multi_base_blacklist then pairs |
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
384 |
else add_single_names ((a, ths), pairs); |
21224 | 385 |
|
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
386 |
fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a; |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
387 |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
388 |
(*The single theorems go BEFORE the multiple ones. Blacklist is applied to all.*) |
21588 | 389 |
fun name_thm_pairs ctxt = |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
390 |
let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt) |
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
391 |
val ht = mk_clause_table 900 (*ht for blacklisted theorems*) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
392 |
fun blacklisted x = run_blacklist_filter andalso isSome (Polyhash.peek ht x) |
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
393 |
in |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
394 |
app (fn th => ignore (Polyhash.peekInsert ht (th,()))) (ResBlacklist.get ctxt); |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
395 |
filter (not o blacklisted o #2) |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
396 |
(foldl add_single_names (foldl add_multi_names [] mults) singles) |
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24215
diff
changeset
|
397 |
end; |
21224 | 398 |
|
26928 | 399 |
fun check_named ("",th) = (warning ("No name for theorem " ^ Display.string_of_thm th); false) |
21224 | 400 |
| check_named (_,th) = true; |
19894 | 401 |
|
26928 | 402 |
fun display_thm (name,th) = Output.debug (fn () => name ^ ": " ^ Display.string_of_thm th); |
22193 | 403 |
|
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
404 |
(* get lemmas from claset, simpset, atpset and extra supplied rules *) |
21588 | 405 |
fun get_clasimp_atp_lemmas ctxt user_thms = |
19894 | 406 |
let val included_thms = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
407 |
if include_all |
21588 | 408 |
then (tap (fn ths => Output.debug |
22130 | 409 |
(fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems"))) |
21588 | 410 |
(name_thm_pairs ctxt)) |
411 |
else |
|
412 |
let val claset_thms = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
413 |
if include_claset then ResAxioms.claset_rules_of ctxt |
21588 | 414 |
else [] |
415 |
val simpset_thms = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
416 |
if include_simpset then ResAxioms.simpset_rules_of ctxt |
21588 | 417 |
else [] |
418 |
val atpset_thms = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
419 |
if include_atpset then ResAxioms.atpset_rules_of ctxt |
21588 | 420 |
else [] |
22193 | 421 |
val _ = (Output.debug (fn () => "ATP theorems: "); app display_thm atpset_thms) |
21588 | 422 |
in claset_thms @ simpset_thms @ atpset_thms end |
423 |
val user_rules = filter check_named |
|
22382
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
paulson
parents:
22217
diff
changeset
|
424 |
(map ResAxioms.pairname |
24215 | 425 |
(if null user_thms then whitelist else user_thms)) |
19894 | 426 |
in |
21224 | 427 |
(filter check_named included_thms, user_rules) |
19894 | 428 |
end; |
19768
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents:
19746
diff
changeset
|
429 |
|
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
430 |
(***************************************************************) |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
431 |
(* Type Classes Present in the Axiom or Conjecture Clauses *) |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
432 |
(***************************************************************) |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
433 |
|
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
434 |
fun add_classes (sorts, cset) = foldl setinsert cset (List.concat sorts); |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
435 |
|
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
436 |
(*Remove this trivial type class*) |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
437 |
fun delete_type cset = Symtab.delete_safe "HOL.type" cset; |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
438 |
|
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
439 |
fun tvar_classes_of_terms ts = |
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29267
diff
changeset
|
440 |
let val sorts_list = map (map #2 o OldTerm.term_tvars) ts |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
441 |
in Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list)) end; |
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
442 |
|
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
443 |
fun tfree_classes_of_terms ts = |
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29267
diff
changeset
|
444 |
let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
445 |
in Symtab.keys (delete_type (foldl add_classes Symtab.empty sorts_list)) end; |
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
446 |
|
21373
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
447 |
(*fold type constructors*) |
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
448 |
fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x)) |
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
449 |
| fold_type_consts f T x = x; |
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
450 |
|
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
451 |
val add_type_consts_in_type = fold_type_consts setinsert; |
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
452 |
|
21397
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
453 |
(*Type constructors used to instantiate overloaded constants are the only ones needed.*) |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
454 |
fun add_type_consts_in_term thy = |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
455 |
let val const_typargs = Sign.const_typargs thy |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
456 |
fun add_tcs (Const cT) x = fold add_type_consts_in_type (const_typargs cT) x |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
457 |
| add_tcs (Abs (_, T, u)) x = add_tcs u x |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
458 |
| add_tcs (t $ u) x = add_tcs t (add_tcs u x) |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
459 |
| add_tcs _ x = x |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
460 |
in add_tcs end |
21373
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
461 |
|
21397
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
462 |
fun type_consts_of_terms thy ts = |
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
paulson
parents:
21373
diff
changeset
|
463 |
Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty); |
21373
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
464 |
|
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
465 |
|
19194 | 466 |
(***************************************************************) |
467 |
(* ATP invocation methods setup *) |
|
468 |
(***************************************************************) |
|
469 |
||
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
470 |
(*Ensures that no higher-order theorems "leak out"*) |
24958 | 471 |
fun restrict_to_logic thy true cls = filter (Meson.is_fol_term thy o prop_of o fst) cls |
472 |
| restrict_to_logic thy false cls = cls; |
|
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
473 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
474 |
(**** Predicates to detect unwanted clauses (prolific or likely to cause unsoundness) ****) |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
475 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
476 |
(** Too general means, positive equality literal with a variable X as one operand, |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
477 |
when X does not occur properly in the other operand. This rules out clearly |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
478 |
inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **) |
21588 | 479 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
480 |
fun occurs ix = |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
481 |
let fun occ(Var (jx,_)) = (ix=jx) |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
482 |
| occ(t1$t2) = occ t1 orelse occ t2 |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
483 |
| occ(Abs(_,_,t)) = occ t |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
484 |
| occ _ = false |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
485 |
in occ end; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
486 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
487 |
fun is_recordtype T = not (null (RecordPackage.dest_recTs T)); |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
488 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
489 |
(*Unwanted equalities include |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
490 |
(1) those between a variable that does not properly occur in the second operand, |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
491 |
(2) those between a variable and a record, since these seem to be prolific "cases" thms |
21588 | 492 |
*) |
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
493 |
fun too_general_eqterms (Var (ix,T), t) = not (occurs ix t) orelse is_recordtype T |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
494 |
| too_general_eqterms _ = false; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
495 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
496 |
fun too_general_equality (Const ("op =", _) $ x $ y) = |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
497 |
too_general_eqterms (x,y) orelse too_general_eqterms(y,x) |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
498 |
| too_general_equality _ = false; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
499 |
|
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
500 |
(* tautologous? *) |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
501 |
fun is_taut (Const ("Trueprop", _) $ Const ("True", _)) = true |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
502 |
| is_taut _ = false; |
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
503 |
|
29267 | 504 |
fun has_typed_var tycons = exists_subterm |
505 |
(fn Var (_, Type (a, _)) => member (op =) tycons a | _ => false); |
|
21431
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
506 |
|
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
507 |
(*Clauses are forbidden to contain variables of these types. The typical reason is that |
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
508 |
they lead to unsoundness. Note that "unit" satisfies numerous equations like ?X=(). |
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
509 |
The resulting clause will have no type constraint, yielding false proofs. Even "bool" |
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
510 |
leads to many unsound proofs, though (obviously) only for higher-order problems.*) |
24215 | 511 |
val unwanted_types = ["Product_Type.unit","bool"]; |
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
512 |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
513 |
fun unwanted t = |
24958 | 514 |
is_taut t orelse has_typed_var unwanted_types t orelse |
515 |
forall too_general_equality (HOLogic.disjuncts (strip_Trueprop t)); |
|
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
516 |
|
21431
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
517 |
(*Clauses containing variables of type "unit" or "bool" are unlikely to be useful and |
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
518 |
likely to lead to unsound proofs.*) |
22217
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
paulson
parents:
22193
diff
changeset
|
519 |
fun remove_unwanted_clauses cls = filter (not o unwanted o prop_of o fst) cls; |
21431
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
520 |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
521 |
(* TODO: problem file for *one* subgoal would be sufficient *) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
522 |
(*Write out problem files for each subgoal. |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
523 |
Argument probfile generates filenames from subgoal-number |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
524 |
Arguments max_new and theory_const are booleans controlling relevance_filter |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
525 |
(necessary for different provers) |
30151 | 526 |
*) |
527 |
fun write_problem_files dfg max_new theory_const probfile (ctxt, chain_ths, th) = |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
528 |
let val goals = Thm.prems_of th |
17717 | 529 |
val thy = ProofContext.theory_of ctxt |
20457
85414caac94a
refinements to conversion into clause form, esp for the HO case
paulson
parents:
20446
diff
changeset
|
530 |
fun get_neg_subgoals [] _ = [] |
22865 | 531 |
| get_neg_subgoals (gl::gls) n = #1 (ResAxioms.neg_conjecture_clauses th n) :: |
21999
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
paulson
parents:
21980
diff
changeset
|
532 |
get_neg_subgoals gls (n+1) |
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
533 |
val goal_cls = get_neg_subgoals goals 1 |
25243 | 534 |
handle THM ("assume: variables", _, _) => |
535 |
error "Sledgehammer: Goal contains type variables (TVars)" |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
536 |
val isFO = case linkup_logic_mode of |
24958 | 537 |
Auto => forall (Meson.is_fol_term thy) (map prop_of (List.concat goal_cls)) |
538 |
| Fol => true |
|
539 |
| Hol => false |
|
24546 | 540 |
val (included_thms,white_thms) = get_clasimp_atp_lemmas ctxt chain_ths |
27178
c8ddb3000743
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
26928
diff
changeset
|
541 |
val included_cls = included_thms |> ResAxioms.cnf_rules_pairs thy |> make_unique |
24958 | 542 |
|> restrict_to_logic thy isFO |
21470
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents:
21431
diff
changeset
|
543 |
|> remove_unwanted_clauses |
27178
c8ddb3000743
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
26928
diff
changeset
|
544 |
val white_cls = ResAxioms.cnf_rules_pairs thy white_thms |
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
545 |
(*clauses relevant to goal gl*) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
546 |
val axcls_list = map (fn ngcls => white_cls @ relevance_filter max_new theory_const thy included_cls (map prop_of ngcls)) goal_cls |
30151 | 547 |
val writer = if dfg then ResHolClause.dfg_write_file else ResHolClause.tptp_write_file |
20526
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
paulson
parents:
20457
diff
changeset
|
548 |
fun write_all [] [] _ = [] |
21588 | 549 |
| write_all (ccls::ccls_list) (axcls::axcls_list) k = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
28065
diff
changeset
|
550 |
let val fname = File.platform_path (probfile k) |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
551 |
val axcls = make_unique axcls |
26928 | 552 |
val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) ccls |
20868
724ab9896068
Improved detection of identical clauses, also in the conjecture
paulson
parents:
20854
diff
changeset
|
553 |
val ccls = subtract_cls ccls axcls |
26928 | 554 |
val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) ccls |
21431
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
555 |
val ccltms = map prop_of ccls |
21373
18f519614978
Arity clauses are now produced only for types and type classes actually used.
paulson
parents:
21311
diff
changeset
|
556 |
and axtms = map (prop_of o #1) axcls |
21431
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
557 |
val subs = tfree_classes_of_terms ccltms |
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
558 |
and supers = tvar_classes_of_terms axtms |
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents:
21397
diff
changeset
|
559 |
and tycons = type_consts_of_terms thy (ccltms@axtms) |
21290
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents:
21243
diff
changeset
|
560 |
(*TFrees in conjecture clauses; TVars in axiom clauses*) |
30151 | 561 |
val (supers',arity_clauses) = ResClause.make_arity_clauses_dfg dfg thy tycons supers |
22645 | 562 |
val classrel_clauses = ResClause.make_classrel_clauses thy subs supers' |
24958 | 563 |
val clnames = writer thy isFO ccls fname (axcls,classrel_clauses,arity_clauses) [] |
21888 | 564 |
val thm_names = Vector.fromList clnames |
20870 | 565 |
in (thm_names,fname) :: write_all ccls_list axcls_list (k+1) end |
566 |
val (thm_names_list, filenames) = ListPair.unzip (write_all goal_cls axcls_list 1) |
|
19194 | 567 |
in |
20870 | 568 |
(filenames, thm_names_list) |
19194 | 569 |
end; |
15644 | 570 |
|
15347 | 571 |
end; |