src/HOL/Mirabelle/Tools/sledgehammer_tactics.ML
changeset 41749 1e3a8807ebd4
parent 41741 839d1488045f
child 41959 b460124855b8
--- a/src/HOL/Mirabelle/Tools/sledgehammer_tactics.ML	Thu Feb 10 17:17:31 2011 +0100
+++ b/src/HOL/Mirabelle/Tools/sledgehammer_tactics.ML	Thu Feb 10 17:18:52 2011 +0100
@@ -39,13 +39,6 @@
           relevance_thresholds
           (the_default default_max_relevant max_relevant) is_built_in_const
           relevance_fudge relevance_override chained_ths hyp_ts concl_t
-    (* Check for constants other than the built-in HOL constants. If none of
-       them appear (as should be the case for TPTP problems, unless "auto" or
-       "simp" already did its "magic"), we can skip the relevance filter. *)
-    val pure_goal =
-      not (exists_Const (fn (s, _) => String.isSubstring "." s andalso
-                                      not (String.isSubstring "HOL" s))
-                        (prop_of goal))
     val problem =
       {state = Proof.init ctxt, goal = goal, subgoal = i, subgoal_count = n,
        facts = map Sledgehammer_Provers.Untranslated_Fact facts,