1 (* Title: HOL/Sledgehammer/sledgehammer_hol_clause.ML
2 Author: Jia Meng, NICTA
4 FOL clauses translated from HOL formulae.
7 signature SLEDGEHAMMER_HOL_CLAUSE =
9 type kind = Sledgehammer_FOL_Clause.kind
10 type fol_type = Sledgehammer_FOL_Clause.fol_type
11 type classrel_clause = Sledgehammer_FOL_Clause.classrel_clause
12 type arity_clause = Sledgehammer_FOL_Clause.arity_clause
13 type axiom_name = string
15 type hol_clause_id = int
18 CombConst of string * fol_type * fol_type list (* Const and Free *) |
19 CombVar of string * fol_type |
20 CombApp of combterm * combterm
21 datatype literal = Literal of polarity * combterm
23 HOLClause of {clause_id: hol_clause_id, axiom_name: axiom_name, th: thm,
24 kind: kind, literals: literal list, ctypes_sorts: typ list}
26 val type_of_combterm : combterm -> fol_type
27 val strip_combterm_comb : combterm -> combterm * combterm list
28 val literals_of_term : theory -> term -> literal list * typ list
30 val make_conjecture_clauses : bool -> theory -> thm list -> hol_clause list
31 val make_axiom_clauses : bool -> theory ->
32 (thm * (axiom_name * hol_clause_id)) list -> (axiom_name * hol_clause) list
33 val get_helper_clauses : bool -> theory -> bool ->
34 hol_clause list * (thm * (axiom_name * hol_clause_id)) list * string list ->
36 val write_tptp_file : bool -> Path.T ->
37 hol_clause list * hol_clause list * hol_clause list * hol_clause list *
38 classrel_clause list * arity_clause list ->
40 val write_dfg_file : bool -> Path.T ->
41 hol_clause list * hol_clause list * hol_clause list * hol_clause list *
42 classrel_clause list * arity_clause list -> int * int
45 structure Sledgehammer_HOL_Clause : SLEDGEHAMMER_HOL_CLAUSE =
48 open Sledgehammer_Util
49 open Sledgehammer_FOL_Clause
50 open Sledgehammer_Fact_Preprocessor
52 (* Parameter "full_types" below indicates that full type information is to be
55 (* If true, each function will be directly applied to as many arguments as
56 possible, avoiding use of the "apply" operator. Use of hBOOL is also
58 val minimize_applies = true;
60 fun min_arity_of const_min_arity c = the_default 0 (Symtab.lookup const_min_arity c);
62 (*True if the constant ever appears outside of the top-level position in literals.
63 If false, the constant always receives all of its arguments and is used as a predicate.*)
64 fun needs_hBOOL const_needs_hBOOL c =
65 not minimize_applies orelse
66 the_default false (Symtab.lookup const_needs_hBOOL c);
69 (******************************************************)
70 (* data types for typed combinator expressions *)
71 (******************************************************)
73 type axiom_name = string;
75 type hol_clause_id = int;
78 CombConst of string * fol_type * fol_type list (* Const and Free *) |
79 CombVar of string * fol_type |
80 CombApp of combterm * combterm
82 datatype literal = Literal of polarity * combterm;
85 HOLClause of {clause_id: hol_clause_id, axiom_name: axiom_name, th: thm,
86 kind: kind, literals: literal list, ctypes_sorts: typ list};
89 (*********************************************************************)
90 (* convert a clause with type Term.term to a clause with type clause *)
91 (*********************************************************************)
93 fun isFalse (Literal(pol, CombConst(c,_,_))) =
94 (pol andalso c = "c_False") orelse (not pol andalso c = "c_True")
97 fun isTrue (Literal (pol, CombConst(c,_,_))) =
98 (pol andalso c = "c_True") orelse
99 (not pol andalso c = "c_False")
102 fun isTaut (HOLClause {literals,...}) = exists isTrue literals;
104 fun type_of dfg (Type (a, Ts)) =
105 let val (folTypes,ts) = types_of dfg Ts
106 in (Comp(make_fixed_type_const dfg a, folTypes), ts) end
107 | type_of _ (tp as TFree (a, _)) = (AtomF (make_fixed_type_var a), [tp])
108 | type_of _ (tp as TVar (v, _)) = (AtomV (make_schematic_type_var v), [tp])
109 and types_of dfg Ts =
110 let val (folTyps,ts) = ListPair.unzip (map (type_of dfg) Ts)
111 in (folTyps, union_all ts) end;
113 (* same as above, but no gathering of sort information *)
114 fun simp_type_of dfg (Type (a, Ts)) =
115 Comp(make_fixed_type_const dfg a, map (simp_type_of dfg) Ts)
116 | simp_type_of _ (TFree (a, _)) = AtomF (make_fixed_type_var a)
117 | simp_type_of _ (TVar (v, _)) = AtomV (make_schematic_type_var v);
120 fun const_type_of dfg thy (c,t) =
121 let val (tp,ts) = type_of dfg t
122 in (tp, ts, map (simp_type_of dfg) (Sign.const_typargs thy (c,t))) end;
124 (* convert a Term.term (with combinators) into a combterm, also accummulate sort info *)
125 fun combterm_of dfg thy (Const(c,t)) =
126 let val (tp,ts,tvar_list) = const_type_of dfg thy (c,t)
127 val c' = CombConst(make_fixed_const dfg c, tp, tvar_list)
129 | combterm_of dfg _ (Free(v,t)) =
130 let val (tp,ts) = type_of dfg t
131 val v' = CombConst(make_fixed_var v, tp, [])
133 | combterm_of dfg _ (Var(v,t)) =
134 let val (tp,ts) = type_of dfg t
135 val v' = CombVar(make_schematic_var v,tp)
137 | combterm_of dfg thy (P $ Q) =
138 let val (P',tsP) = combterm_of dfg thy P
139 val (Q',tsQ) = combterm_of dfg thy Q
140 in (CombApp(P',Q'), union (op =) tsP tsQ) end
141 | combterm_of _ _ (t as Abs _) = raise CLAUSE ("HOL CLAUSE", t);
143 fun predicate_of dfg thy ((@{const Not} $ P), polarity) = predicate_of dfg thy (P, not polarity)
144 | predicate_of dfg thy (t,polarity) = (combterm_of dfg thy (Envir.eta_contract t), polarity);
146 fun literals_of_term1 dfg thy args (@{const Trueprop} $ P) = literals_of_term1 dfg thy args P
147 | literals_of_term1 dfg thy args (@{const "op |"} $ P $ Q) =
148 literals_of_term1 dfg thy (literals_of_term1 dfg thy args P) Q
149 | literals_of_term1 dfg thy (lits,ts) P =
150 let val ((pred,ts'),pol) = predicate_of dfg thy (P,true)
152 (Literal(pol,pred)::lits, union (op =) ts ts')
155 fun literals_of_term_dfg dfg thy P = literals_of_term1 dfg thy ([],[]) P;
156 val literals_of_term = literals_of_term_dfg false;
158 (* Trivial problem, which resolution cannot handle (empty clause) *)
161 (* making axiom and conjecture clauses *)
162 fun make_clause dfg thy (clause_id, axiom_name, kind, th) =
163 let val (lits,ctypes_sorts) = literals_of_term_dfg dfg thy (prop_of th)
165 if forall isFalse lits then
168 HOLClause {clause_id = clause_id, axiom_name = axiom_name, th = th,
169 kind = kind, literals = lits, ctypes_sorts = ctypes_sorts}
173 fun add_axiom_clause dfg thy ((th,(name,id)), pairs) =
174 let val cls = make_clause dfg thy (id, name, Axiom, th)
176 if isTaut cls then pairs else (name,cls)::pairs
179 fun make_axiom_clauses dfg thy = List.foldl (add_axiom_clause dfg thy) [];
181 fun make_conjecture_clauses_aux _ _ _ [] = []
182 | make_conjecture_clauses_aux dfg thy n (th::ths) =
183 make_clause dfg thy (n,"conjecture", Conjecture, th) ::
184 make_conjecture_clauses_aux dfg thy (n+1) ths;
186 fun make_conjecture_clauses dfg thy = make_conjecture_clauses_aux dfg thy 0;
189 (**********************************************************************)
190 (* convert clause into ATP specific formats: *)
191 (* TPTP used by Vampire and E *)
192 (* DFG used by SPASS *)
193 (**********************************************************************)
195 (*Result of a function type; no need to check that the argument type matches.*)
196 fun result_type (Comp ("tc_fun", [_, tp2])) = tp2
197 | result_type _ = error "result_type"
199 fun type_of_combterm (CombConst (_, tp, _)) = tp
200 | type_of_combterm (CombVar (_, tp)) = tp
201 | type_of_combterm (CombApp (t1, _)) = result_type (type_of_combterm t1);
203 (*gets the head of a combinator application, along with the list of arguments*)
204 fun strip_combterm_comb u =
205 let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts)
209 val type_wrapper = "ti";
211 fun head_needs_hBOOL const_needs_hBOOL (CombConst(c,_,_)) = needs_hBOOL const_needs_hBOOL c
212 | head_needs_hBOOL _ _ = true;
214 fun wrap_type full_types (s, tp) =
215 if full_types then type_wrapper ^ paren_pack [s, string_of_fol_type tp] else s;
217 fun apply ss = "hAPP" ^ paren_pack ss;
219 fun rev_apply (v, []) = v
220 | rev_apply (v, arg::args) = apply [rev_apply (v, args), arg];
222 fun string_apply (v, args) = rev_apply (v, rev args);
224 (*Apply an operator to the argument strings, using either the "apply" operator or
225 direct function application.*)
226 fun string_of_applic full_types cma (CombConst (c, _, tvars), args) =
227 let val c = if c = "equal" then "c_fequal" else c
228 val nargs = min_arity_of cma c
229 val args1 = List.take(args, nargs)
230 handle Subscript => error ("string_of_applic: " ^ c ^ " has arity " ^
231 Int.toString nargs ^ " but is applied to " ^
232 space_implode ", " args)
233 val args2 = List.drop(args, nargs)
234 val targs = if full_types then [] else map string_of_fol_type tvars
236 string_apply (c ^ paren_pack (args1@targs), args2)
238 | string_of_applic _ _ (CombVar (v, _), args) = string_apply (v, args)
239 | string_of_applic _ _ _ = error "string_of_applic";
241 fun wrap_type_if full_types cnh (head, s, tp) =
242 if head_needs_hBOOL cnh head then wrap_type full_types (s, tp) else s;
244 fun string_of_combterm (params as (full_types, cma, cnh)) t =
245 let val (head, args) = strip_combterm_comb t
246 in wrap_type_if full_types cnh (head,
247 string_of_applic full_types cma
248 (head, map (string_of_combterm (params)) args),
252 (*Boolean-valued terms are here converted to literals.*)
253 fun boolify params t =
254 "hBOOL" ^ paren_pack [string_of_combterm params t];
256 fun string_of_predicate (params as (_,_,cnh)) t =
258 (CombApp(CombApp(CombConst("equal",_,_), t1), t2)) =>
259 (*DFG only: new TPTP prefers infix equality*)
260 ("equal" ^ paren_pack [string_of_combterm params t1, string_of_combterm params t2])
262 case #1 (strip_combterm_comb t) of
263 CombConst(c,_,_) => if needs_hBOOL cnh c then boolify params t else string_of_combterm params t
264 | _ => boolify params t;
267 (*** tptp format ***)
269 fun tptp_of_equality params pol (t1,t2) =
270 let val eqop = if pol then " = " else " != "
271 in string_of_combterm params t1 ^ eqop ^ string_of_combterm params t2 end;
273 fun tptp_literal params (Literal(pol, CombApp(CombApp(CombConst("equal", _, _), t1), t2))) =
274 tptp_of_equality params pol (t1,t2)
275 | tptp_literal params (Literal(pol,pred)) =
276 tptp_sign pol (string_of_predicate params pred);
278 (*Given a clause, returns its literals paired with a list of literals concerning TFrees;
279 the latter should only occur in conjecture clauses.*)
280 fun tptp_type_lits params pos (HOLClause {literals, ctypes_sorts, ...}) =
281 (map (tptp_literal params) literals,
282 map (tptp_of_typeLit pos) (add_typs ctypes_sorts));
284 fun clause2tptp params (cls as HOLClause {axiom_name, clause_id, kind, ...}) =
285 let val (lits,tylits) = tptp_type_lits params (kind = Conjecture) cls
287 (gen_tptp_cls (clause_id, axiom_name, kind, lits, tylits), tylits)
293 fun dfg_literal params (Literal(pol,pred)) = dfg_sign pol (string_of_predicate params pred);
295 fun dfg_type_lits params pos (HOLClause {literals, ctypes_sorts, ...}) =
296 (map (dfg_literal params) literals,
297 map (dfg_of_typeLit pos) (add_typs ctypes_sorts));
299 fun get_uvars (CombConst _) vars = vars
300 | get_uvars (CombVar(v,_)) vars = (v::vars)
301 | get_uvars (CombApp(P,Q)) vars = get_uvars P (get_uvars Q vars);
303 fun get_uvars_l (Literal(_,c)) = get_uvars c [];
305 fun dfg_vars (HOLClause {literals,...}) = union_all (map get_uvars_l literals);
307 fun clause2dfg params (cls as HOLClause {axiom_name, clause_id, kind,
308 ctypes_sorts, ...}) =
309 let val (lits,tylits) = dfg_type_lits params (kind = Conjecture) cls
310 val vars = dfg_vars cls
311 val tvars = get_tvar_strs ctypes_sorts
313 (gen_dfg_cls (clause_id, axiom_name, kind, lits, tylits, tvars@vars), tylits)
317 (** For DFG format: accumulate function and predicate declarations **)
319 fun addtypes tvars tab = List.foldl add_foltype_funcs tab tvars;
321 fun add_decls (full_types, cma, cnh) (CombConst (c, _, tvars), (funcs, preds)) =
322 if c = "equal" then (addtypes tvars funcs, preds)
324 let val arity = min_arity_of cma c
325 val ntys = if not full_types then length tvars else 0
326 val addit = Symtab.update(c, arity+ntys)
328 if needs_hBOOL cnh c then (addtypes tvars (addit funcs), preds)
329 else (addtypes tvars funcs, addit preds)
331 | add_decls _ (CombVar(_,ctp), (funcs,preds)) =
332 (add_foltype_funcs (ctp,funcs), preds)
333 | add_decls params (CombApp(P,Q),decls) = add_decls params (P,add_decls params (Q,decls));
335 fun add_literal_decls params (Literal (_,c), decls) = add_decls params (c,decls);
337 fun add_clause_decls params (HOLClause {literals, ...}, decls) =
338 List.foldl (add_literal_decls params) decls literals
339 handle Symtab.DUP a => error ("function " ^ a ^ " has multiple arities")
341 fun decls_of_clauses params clauses arity_clauses =
342 let val init_functab = Symtab.update (type_wrapper,2) (Symtab.update ("hAPP",2) init_functab)
343 val init_predtab = Symtab.update ("hBOOL",1) Symtab.empty
344 val (functab,predtab) = (List.foldl (add_clause_decls params) (init_functab, init_predtab) clauses)
346 (Symtab.dest (List.foldl add_arity_clause_funcs functab arity_clauses),
350 fun add_clause_preds (HOLClause {ctypes_sorts, ...}, preds) =
351 List.foldl add_type_sort_preds preds ctypes_sorts
352 handle Symtab.DUP a => error ("predicate " ^ a ^ " has multiple arities")
354 (*Higher-order clauses have only the predicates hBOOL and type classes.*)
355 fun preds_of_clauses clauses clsrel_clauses arity_clauses =
357 (List.foldl add_classrel_clause_preds
358 (List.foldl add_arity_clause_preds
359 (List.foldl add_clause_preds Symtab.empty clauses)
364 (**********************************************************************)
365 (* write clauses to files *)
366 (**********************************************************************)
369 Symtab.make [("c_COMBI", 0), ("c_COMBK", 0), ("c_COMBB", 0), ("c_COMBC", 0),
372 fun count_combterm (CombConst (c, _, _), ct) =
373 (case Symtab.lookup ct c of NONE => ct (*no counter*)
374 | SOME n => Symtab.update (c,n+1) ct)
375 | count_combterm (CombVar _, ct) = ct
376 | count_combterm (CombApp(t1,t2), ct) = count_combterm(t1, count_combterm(t2, ct));
378 fun count_literal (Literal(_,t), ct) = count_combterm(t,ct);
380 fun count_clause (HOLClause {literals, ...}, ct) =
381 List.foldl count_literal ct literals;
383 fun count_user_clause user_lemmas (HOLClause {axiom_name, literals, ...}, ct) =
384 if axiom_name mem_string user_lemmas then List.foldl count_literal ct literals
387 fun cnf_helper_thms thy = cnf_rules_pairs thy o map pairname
389 fun get_helper_clauses dfg thy isFO (conjectures, axcls, user_lemmas) =
394 val axclauses = map #2 (make_axiom_clauses dfg thy axcls)
395 val ct0 = List.foldl count_clause init_counters conjectures
396 val ct = List.foldl (count_user_clause user_lemmas) ct0 axclauses
397 fun needed c = the (Symtab.lookup ct c) > 0
398 val IK = if needed "c_COMBI" orelse needed "c_COMBK"
399 then cnf_helper_thms thy [@{thm COMBI_def}, @{thm COMBK_def}]
401 val BC = if needed "c_COMBB" orelse needed "c_COMBC"
402 then cnf_helper_thms thy [@{thm COMBB_def}, @{thm COMBC_def}]
404 val S = if needed "c_COMBS" then cnf_helper_thms thy [@{thm COMBS_def}]
406 val other = cnf_helper_thms thy [@{thm fequal_imp_equal},
407 @{thm equal_imp_fequal}]
409 map #2 (make_axiom_clauses dfg thy (other @ IK @ BC @ S))
412 (*Find the minimal arity of each function mentioned in the term. Also, note which uses
413 are not at top level, to see if hBOOL is needed.*)
414 fun count_constants_term toplev t (const_min_arity, const_needs_hBOOL) =
415 let val (head, args) = strip_combterm_comb t
417 val (const_min_arity, const_needs_hBOOL) = fold (count_constants_term false) args (const_min_arity, const_needs_hBOOL)
420 CombConst (a,_,_) => (*predicate or function version of "equal"?*)
421 let val a = if a="equal" andalso not toplev then "c_fequal" else a
422 val const_min_arity = Symtab.map_default (a, n) (Integer.min n) const_min_arity
424 if toplev then (const_min_arity, const_needs_hBOOL)
425 else (const_min_arity, Symtab.update (a,true) (const_needs_hBOOL))
427 | _ => (const_min_arity, const_needs_hBOOL)
430 (*A literal is a top-level term*)
431 fun count_constants_lit (Literal (_,t)) (const_min_arity, const_needs_hBOOL) =
432 count_constants_term true t (const_min_arity, const_needs_hBOOL);
434 fun count_constants_clause (HOLClause {literals, ...})
435 (const_min_arity, const_needs_hBOOL) =
436 fold count_constants_lit literals (const_min_arity, const_needs_hBOOL);
438 fun display_arity const_needs_hBOOL (c,n) =
439 trace_msg (fn () => "Constant: " ^ c ^
440 " arity:\t" ^ Int.toString n ^
441 (if needs_hBOOL const_needs_hBOOL c then " needs hBOOL" else ""));
443 fun count_constants (conjectures, _, extra_clauses, helper_clauses, _, _) =
444 if minimize_applies then
445 let val (const_min_arity, const_needs_hBOOL) =
446 fold count_constants_clause conjectures (Symtab.empty, Symtab.empty)
447 |> fold count_constants_clause extra_clauses
448 |> fold count_constants_clause helper_clauses
449 val _ = List.app (display_arity const_needs_hBOOL) (Symtab.dest (const_min_arity))
450 in (const_min_arity, const_needs_hBOOL) end
451 else (Symtab.empty, Symtab.empty);
455 fun write_tptp_file full_types file clauses =
457 fun section _ [] = []
458 | section name ss = "\n% " ^ name ^ plural_s (length ss) ^ "\n" :: ss
459 val (conjectures, axclauses, _, helper_clauses,
460 classrel_clauses, arity_clauses) = clauses
461 val (cma, cnh) = count_constants clauses
462 val params = (full_types, cma, cnh)
463 val (tptp_clss,tfree_litss) = ListPair.unzip (map (clause2tptp params) conjectures)
464 val tfree_clss = map tptp_tfree_clause (List.foldl (uncurry (union (op =))) [] tfree_litss)
465 val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" (Date.fromTimeLocal (Time.now ()))
467 File.write_list file (
468 "% This file was generated by Isabelle (most likely Sledgehammer)\n" ^
469 "% " ^ timestamp ^ "\n" ::
470 section "Relevant fact" (map (#1 o (clause2tptp params)) axclauses) @
471 section "Type variable" tfree_clss @
472 section "Class relationship"
473 (map tptp_classrel_clause classrel_clauses) @
474 section "Arity declaration" (map tptp_arity_clause arity_clauses) @
475 section "Helper fact" (map (#1 o (clause2tptp params)) helper_clauses) @
476 section "Conjecture" tptp_clss)
477 in (length axclauses + 1, length tfree_clss + length tptp_clss)
483 fun write_dfg_file full_types file clauses =
485 val (conjectures, axclauses, _, helper_clauses,
486 classrel_clauses, arity_clauses) = clauses
487 val (cma, cnh) = count_constants clauses
488 val params = (full_types, cma, cnh)
489 val (dfg_clss, tfree_litss) = ListPair.unzip (map (clause2dfg params) conjectures)
490 and probname = Path.implode (Path.base file)
491 val axstrs = map (#1 o (clause2dfg params)) axclauses
492 val tfree_clss = map dfg_tfree_clause (union_all tfree_litss)
493 val helper_clauses_strs = map (#1 o (clause2dfg params)) helper_clauses
494 val (funcs,cl_preds) = decls_of_clauses params (helper_clauses @ conjectures @ axclauses) arity_clauses
495 and ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses
497 File.write_list file (
498 string_of_start probname ::
499 string_of_descrip probname ::
500 string_of_symbols (string_of_funcs funcs)
501 (string_of_preds (cl_preds @ ty_preds)) ::
502 "list_of_clauses(axioms, cnf).\n" ::
504 map dfg_classrel_clause classrel_clauses @
505 map dfg_arity_clause arity_clauses @
506 helper_clauses_strs @
507 ["end_of_list.\n\nlist_of_clauses(conjectures, cnf).\n"] @
511 (*VarWeight=3 helps the HO problems, probably by counteracting the presence of hAPP*)
512 "list_of_settings(SPASS).\n{*\nset_flag(VarWeight, 3).\n*}\nend_of_list.\n\n",
515 in (length axclauses + length classrel_clauses + length arity_clauses +
516 length helper_clauses + 1, length tfree_clss + length dfg_clss)