--- a/src/HOL/Tools/ATP/atp_problem.ML Tue Jan 31 14:39:21 2012 +0100
+++ b/src/HOL/Tools/ATP/atp_problem.ML Tue Jan 31 15:10:03 2012 +0100
@@ -432,6 +432,8 @@
fun is_predicate_type (AFun (_, ty)) = is_predicate_type ty
| is_predicate_type (AType (s, _)) = (s = tptp_bool_type)
| is_predicate_type _ = false
+fun is_nontrivial_predicate_type (AFun (_, ty)) = is_predicate_type ty
+ | is_nontrivial_predicate_type _ = false
fun dfg_string_for_formula flavor info =
let
@@ -495,7 +497,8 @@
|> commas |> enclose "functions [" "]."
val pred_aries =
filt (fn Decl (_, sym, ty) =>
- if is_predicate_type ty then SOME (ary sym ty) else NONE
+ if is_nontrivial_predicate_type ty then SOME (ary sym ty)
+ else NONE
| _ => NONE)
|> commas |> enclose "predicates [" "]."
fun sorts () =
--- a/src/HOL/Tools/ATP/atp_systems.ML Tue Jan 31 14:39:21 2012 +0100
+++ b/src/HOL/Tools/ATP/atp_systems.ML Tue Jan 31 15:10:03 2012 +0100
@@ -360,7 +360,7 @@
known_failures = #known_failures spass_config,
conj_sym_kind = #conj_sym_kind spass_config,
prem_kind = #prem_kind spass_config,
- best_slices = fn ctxt =>
+ best_slices = fn _ =>
(* FUDGE *)
[(0.333, (false, (150, DFG DFG_Sorted, "poly_tags??", combs_and_liftingN,
""))),