src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
author blanchet
Wed, 18 Aug 2010 19:57:12 +0200
changeset 38817 af205f4fd0f3
parent 38816 85aef7587cf1
child 38818 bbb0982656eb
permissions -rw-r--r--
tuning of relevance filter
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
     2     Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
     3     Author:     Jasmin Blanchette, TU Muenchen
     4 *)
     5 
     6 signature SLEDGEHAMMER_FACT_FILTER =
     7 sig
     8   type relevance_override =
     9     {add: Facts.ref list,
    10      del: Facts.ref list,
    11      only: bool}
    12 
    13   val trace : bool Unsynchronized.ref
    14   val chained_prefix : string
    15   val name_thms_pair_from_ref :
    16     Proof.context -> thm list -> Facts.ref -> string * thm list
    17   val relevant_facts :
    18     bool -> real -> real -> bool -> int -> bool -> relevance_override
    19     -> Proof.context * (thm list * 'a) -> term list -> term
    20     -> (string * thm) list
    21 end;
    22 
    23 structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
    24 struct
    25 
    26 val trace = Unsynchronized.ref false
    27 fun trace_msg msg = if !trace then tracing (msg ()) else ()
    28 
    29 val respect_no_atp = true
    30 
    31 type relevance_override =
    32   {add: Facts.ref list,
    33    del: Facts.ref list,
    34    only: bool}
    35 
    36 val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
    37 (* Used to label theorems chained into the goal. *)
    38 val chained_prefix = sledgehammer_prefix ^ "chained_"
    39 
    40 fun name_thms_pair_from_ref ctxt chained_ths xref =
    41   let
    42     val ths = ProofContext.get_fact ctxt xref
    43     val name = Facts.string_of_ref xref
    44                |> forall (member Thm.eq_thm chained_ths) ths
    45                   ? prefix chained_prefix
    46   in (name, ths) end
    47 
    48 
    49 (***************************************************************)
    50 (* Relevance Filtering                                         *)
    51 (***************************************************************)
    52 
    53 fun strip_Trueprop_and_all (@{const Trueprop} $ t) =
    54     strip_Trueprop_and_all t
    55   | strip_Trueprop_and_all (Const (@{const_name all}, _) $ Abs (_, _, t)) =
    56     strip_Trueprop_and_all t
    57   | strip_Trueprop_and_all (Const (@{const_name All}, _) $ Abs (_, _, t)) =
    58     strip_Trueprop_and_all t
    59   | strip_Trueprop_and_all t = t
    60 
    61 (*** constants with types ***)
    62 
    63 (*An abstraction of Isabelle types*)
    64 datatype const_typ =  CTVar | CType of string * const_typ list
    65 
    66 (*Is the second type an instance of the first one?*)
    67 fun match_type (CType(con1,args1)) (CType(con2,args2)) =
    68       con1=con2 andalso match_types args1 args2
    69   | match_type CTVar _ = true
    70   | match_type _ CTVar = false
    71 and match_types [] [] = true
    72   | match_types (a1::as1) (a2::as2) = match_type a1 a2 andalso match_types as1 as2;
    73 
    74 (*Is there a unifiable constant?*)
    75 fun uni_mem goal_const_tab (c, c_typ) =
    76   exists (match_types c_typ) (these (Symtab.lookup goal_const_tab c))
    77 
    78 (*Maps a "real" type to a const_typ*)
    79 fun const_typ_of (Type (c,typs)) = CType (c, map const_typ_of typs)
    80   | const_typ_of (TFree _) = CTVar
    81   | const_typ_of (TVar _) = CTVar
    82 
    83 (*Pairs a constant with the list of its type instantiations (using const_typ)*)
    84 fun const_with_typ thy (c,typ) =
    85     let val tvars = Sign.const_typargs thy (c,typ)
    86     in (c, map const_typ_of tvars) end
    87     handle TYPE _ => (c,[]);   (*Variable (locale constant): monomorphic*)
    88 
    89 (*Add a const/type pair to the table, but a [] entry means a standard connective,
    90   which we ignore.*)
    91 fun add_const_type_to_table (c, ctyps) =
    92   Symtab.map_default (c, [ctyps])
    93                      (fn [] => [] | ctypss => insert (op =) ctyps ctypss)
    94 
    95 val fresh_prefix = "Sledgehammer.Fresh."
    96 val flip = Option.map not
    97 (* These are typically simplified away by "Meson.presimplify". *)
    98 val boring_consts = [@{const_name If}, @{const_name Let}]
    99 
   100 fun get_consts_typs thy pos ts =
   101   let
   102     (* We include free variables, as well as constants, to handle locales. For
   103        each quantifiers that must necessarily be skolemized by the ATP, we
   104        introduce a fresh constant to simulate the effect of Skolemization. *)
   105     fun do_term t =
   106       case t of
   107         Const (@{const_name Let}, _) => I
   108       | Const x => add_const_type_to_table (const_with_typ thy x)
   109       | Free x => add_const_type_to_table (const_with_typ thy x)
   110       | t1 $ t2 => do_term t1 #> do_term t2
   111       | Abs (_, _, t) =>
   112         (* Abstractions lead to combinators, so we add a penalty for them. *)
   113         add_const_type_to_table (gensym fresh_prefix, [])
   114         #> do_term t
   115       | _ => I
   116     fun do_quantifier will_surely_be_skolemized body_t =
   117       do_formula pos body_t
   118       #> (if will_surely_be_skolemized then
   119             add_const_type_to_table (gensym fresh_prefix, [])
   120           else
   121             I)
   122     and do_term_or_formula T =
   123       if T = @{typ bool} orelse T = @{typ prop} then do_formula NONE
   124       else do_term
   125     and do_formula pos t =
   126       case t of
   127         Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
   128         do_quantifier (pos = SOME false) body_t
   129       | @{const "==>"} $ t1 $ t2 =>
   130         do_formula (flip pos) t1 #> do_formula pos t2
   131       | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
   132         fold (do_term_or_formula T) [t1, t2]
   133       | @{const Trueprop} $ t1 => do_formula pos t1
   134       | @{const Not} $ t1 => do_formula (flip pos) t1
   135       | Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
   136         do_quantifier (pos = SOME false) body_t
   137       | Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
   138         do_quantifier (pos = SOME true) body_t
   139       | @{const "op &"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
   140       | @{const "op |"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
   141       | @{const "op -->"} $ t1 $ t2 =>
   142         do_formula (flip pos) t1 #> do_formula pos t2
   143       | Const (@{const_name "op ="}, Type (_, [T, _])) $ t1 $ t2 =>
   144         fold (do_term_or_formula T) [t1, t2]
   145       | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
   146         $ t1 $ t2 $ t3 =>
   147         do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
   148       | Const (@{const_name Ex1}, _) $ Abs (_, _, body_t) =>
   149         do_quantifier (is_some pos) body_t
   150       | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, body_t) =>
   151         do_quantifier (pos = SOME false)
   152                       (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, body_t))
   153       | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, body_t) =>
   154         do_quantifier (pos = SOME true)
   155                       (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, body_t))
   156       | (t0 as Const (_, @{typ bool})) $ t1 =>
   157         do_term t0 #> do_formula pos t1  (* theory constant *)
   158       | _ => do_term t
   159   in
   160     Symtab.empty |> fold (Symtab.update o rpair []) boring_consts
   161                  |> fold (do_formula pos) ts
   162   end
   163 
   164 (*Inserts a dummy "constant" referring to the theory name, so that relevance
   165   takes the given theory into account.*)
   166 fun theory_const_prop_of theory_relevant th =
   167   if theory_relevant then
   168     let
   169       val name = Context.theory_name (theory_of_thm th)
   170       val t = Const (name ^ ". 1", @{typ bool})
   171     in t $ prop_of th end
   172   else
   173     prop_of th
   174 
   175 (**** Constant / Type Frequencies ****)
   176 
   177 (*A two-dimensional symbol table counts frequencies of constants. It's keyed first by
   178   constant name and second by its list of type instantiations. For the latter, we need
   179   a linear ordering on type const_typ list.*)
   180 
   181 local
   182 
   183 fun cons_nr CTVar = 0
   184   | cons_nr (CType _) = 1;
   185 
   186 in
   187 
   188 fun const_typ_ord TU =
   189   case TU of
   190     (CType (a, Ts), CType (b, Us)) =>
   191       (case fast_string_ord(a,b) of EQUAL => dict_ord const_typ_ord (Ts,Us) | ord => ord)
   192   | (T, U) => int_ord (cons_nr T, cons_nr U);
   193 
   194 end;
   195 
   196 structure CTtab = Table(type key = const_typ list val ord = dict_ord const_typ_ord);
   197 
   198 fun count_axiom_consts theory_relevant thy (_, th) =
   199   let
   200     fun do_const (a, T) =
   201       let val (c, cts) = const_with_typ thy (a,T) in
   202         (* Two-dimensional table update. Constant maps to types maps to
   203            count. *)
   204         CTtab.map_default (cts, 0) (Integer.add 1)
   205         |> Symtab.map_default (c, CTtab.empty)
   206       end
   207     fun do_term (Const x) = do_const x
   208       | do_term (Free x) = do_const x
   209       | do_term (t $ u) = do_term t #> do_term u
   210       | do_term (Abs (_, _, t)) = do_term t
   211       | do_term _ = I
   212   in th |> theory_const_prop_of theory_relevant |> do_term end
   213 
   214 
   215 (**** Actual Filtering Code ****)
   216 
   217 (*The frequency of a constant is the sum of those of all instances of its type.*)
   218 fun const_frequency const_tab (c, cts) =
   219   CTtab.fold (fn (cts', m) => match_types cts cts' ? Integer.add m)
   220              (the (Symtab.lookup const_tab c)
   221               handle Option.Option => raise Fail ("Const: " ^ c)) 0
   222 
   223 (* A surprising number of theorems contain only a few significant constants.
   224    These include all induction rules, and other general theorems. *)
   225 
   226 (* "log" seems best in practice. A constant function of one ignores the constant
   227    frequencies. *)
   228 fun log_weight2 (x:real) = 1.0 + 2.0 / Math.ln (x + 1.0)
   229 
   230 (* Computes a constant's weight, as determined by its frequency. *)
   231 val ct_weight = log_weight2 o real oo const_frequency
   232 
   233 (*Relevant constants are weighted according to frequency,
   234   but irrelevant constants are simply counted. Otherwise, Skolem functions,
   235   which are rare, would harm a formula's chances of being picked.*)
   236 fun formula_weight const_tab gctyps consts_typs =
   237   let
   238     val rel = filter (uni_mem gctyps) consts_typs
   239     val rel_weight = fold (curry Real.+ o ct_weight const_tab) rel 0.0
   240     val res = rel_weight / (rel_weight + real (length consts_typs - length rel))
   241   in if Real.isFinite res then res else 0.0 end
   242 
   243 (*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*)
   244 fun add_expand_pairs (x,ys) xys = List.foldl (fn (y,acc) => (x,y)::acc) xys ys;
   245 
   246 fun consts_typs_of_term thy t =
   247   Symtab.fold add_expand_pairs (get_consts_typs thy (SOME true) [t]) []
   248 
   249 fun pair_consts_typs_axiom theory_relevant thy axiom =
   250   (axiom, axiom |> snd |> theory_const_prop_of theory_relevant
   251                 |> consts_typs_of_term thy)
   252 
   253 exception CONST_OR_FREE of unit
   254 
   255 fun dest_Const_or_Free (Const x) = x
   256   | dest_Const_or_Free (Free x) = x
   257   | dest_Const_or_Free _ = raise CONST_OR_FREE ()
   258 
   259 (*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
   260 fun defines thy thm gctypes =
   261     let val tm = prop_of thm
   262         fun defs lhs rhs =
   263             let val (rator,args) = strip_comb lhs
   264                 val ct = const_with_typ thy (dest_Const_or_Free rator)
   265             in
   266               forall is_Var args andalso uni_mem gctypes ct andalso
   267                 subset (op =) (Term.add_vars rhs [], Term.add_vars lhs [])
   268             end
   269             handle CONST_OR_FREE () => false
   270     in
   271         case tm of
   272           @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ lhs $ rhs) =>
   273             defs lhs rhs
   274         | _ => false
   275     end;
   276 
   277 type annotated_thm = (string * thm) * (string * const_typ list) list
   278 
   279 (*For a reverse sort, putting the largest values first.*)
   280 fun compare_pairs ((_, w1), (_, w2)) = Real.compare (w2, w1)
   281 
   282 (* Limit the number of new facts, to prevent runaway acceptance. *)
   283 fun take_best max_new (newpairs : (annotated_thm * real) list) =
   284   let val nnew = length newpairs in
   285     if nnew <= max_new then
   286       (map #1 newpairs, [])
   287     else
   288       let
   289         val newpairs = sort compare_pairs newpairs
   290         val accepted = List.take (newpairs, max_new)
   291       in
   292         trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^
   293                        ", exceeds the limit of " ^ Int.toString max_new));
   294         trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
   295         trace_msg (fn () => "Actually passed: " ^
   296           space_implode ", " (map (fst o fst o fst) accepted));
   297         (map #1 accepted, map #1 (List.drop (newpairs, max_new)))
   298       end
   299   end;
   300 
   301 fun relevance_filter ctxt relevance_threshold relevance_convergence
   302                      defs_relevant max_new theory_relevant
   303                      ({add, del, ...} : relevance_override) axioms goal_ts =
   304   if relevance_threshold > 1.0 then
   305     []
   306   else if relevance_threshold < 0.0 then
   307     axioms
   308   else
   309     let
   310       val thy = ProofContext.theory_of ctxt
   311       val const_tab = fold (count_axiom_consts theory_relevant thy) axioms
   312                            Symtab.empty
   313       val goal_const_tab = get_consts_typs thy (SOME false) goal_ts
   314       val relevance_threshold = 0.8 * relevance_threshold (* FIXME *)
   315       val _ =
   316         trace_msg (fn () => "Initial constants: " ^
   317                             commas (goal_const_tab
   318                                     |> Symtab.dest
   319                                     |> filter (curry (op <>) [] o snd)
   320                                     |> map fst))
   321       val add_thms = maps (ProofContext.get_fact ctxt) add
   322       val del_thms = maps (ProofContext.get_fact ctxt) del
   323       fun iter threshold rel_const_tab =
   324         let
   325           fun relevant ([], rejects) [] =
   326               (* Nothing was added this iteration: Add "add:" facts. *)
   327               if null add_thms then
   328                 []
   329               else
   330                 map_filter (fn (p as (name, th), _) =>
   331                                if member Thm.eq_thm add_thms th then SOME p
   332                                else NONE) rejects
   333             | relevant (newpairs, rejects) [] =
   334               let
   335                 val (newrels, more_rejects) = take_best max_new newpairs
   336                 val new_consts = maps #2 newrels
   337                 val rel_const_tab =
   338                   rel_const_tab |> fold add_const_type_to_table new_consts
   339                 val threshold =
   340                   threshold + (1.0 - threshold) / relevance_convergence
   341               in
   342                 trace_msg (fn () => "relevant this iteration: " ^
   343                                     Int.toString (length newrels));
   344                 map #1 newrels @ iter threshold rel_const_tab
   345                     (more_rejects @ rejects)
   346               end
   347             | relevant (newrels, rejects)
   348                        ((ax as ((name, th), consts_typs)) :: axs) =
   349               let
   350                 val weight =
   351                   if member Thm.eq_thm del_thms th then 0.0
   352                   else formula_weight const_tab rel_const_tab consts_typs
   353               in
   354                 if weight >= threshold orelse
   355                    (defs_relevant andalso defines thy th rel_const_tab) then
   356                   (trace_msg (fn () =>
   357                        name ^ " passes: " ^ Real.toString weight
   358                        (* ^ " consts: " ^ commas (map fst consts_typs) *));
   359                    relevant ((ax, weight) :: newrels, rejects) axs)
   360                 else
   361                   relevant (newrels, ax :: rejects) axs
   362               end
   363           in
   364             trace_msg (fn () => "relevant_facts, current threshold: " ^
   365                                 Real.toString threshold);
   366             relevant ([], [])
   367           end
   368       val relevant = iter relevance_threshold goal_const_tab
   369                           (map (pair_consts_typs_axiom theory_relevant thy)
   370                                axioms)
   371     in
   372       trace_msg (fn () => "Total relevant: " ^ Int.toString (length relevant));
   373       relevant
   374     end
   375 
   376 (***************************************************************)
   377 (* Retrieving and filtering lemmas                             *)
   378 (***************************************************************)
   379 
   380 (*** retrieve lemmas and filter them ***)
   381 
   382 (*Reject theorems with names like "List.filter.filter_list_def" or
   383   "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
   384 fun is_package_def a =
   385   let val names = Long_Name.explode a
   386   in
   387      length names > 2 andalso
   388      not (hd names = "local") andalso
   389      String.isSuffix "_def" a  orelse  String.isSuffix "_defs" a
   390   end;
   391 
   392 fun make_fact_table xs =
   393   fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty
   394 fun make_unique xs = Termtab.fold (cons o snd) (make_fact_table xs) []
   395 
   396 (* FIXME: put other record thms here, or declare as "no_atp" *)
   397 val multi_base_blacklist =
   398   ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
   399    "split_asm", "cases", "ext_cases"]
   400 
   401 val max_lambda_nesting = 3
   402 
   403 fun term_has_too_many_lambdas max (t1 $ t2) =
   404     exists (term_has_too_many_lambdas max) [t1, t2]
   405   | term_has_too_many_lambdas max (Abs (_, _, t)) =
   406     max = 0 orelse term_has_too_many_lambdas (max - 1) t
   407   | term_has_too_many_lambdas _ _ = false
   408 
   409 fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
   410 
   411 (* Don't count nested lambdas at the level of formulas, since they are
   412    quantifiers. *)
   413 fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
   414     formula_has_too_many_lambdas (T :: Ts) t
   415   | formula_has_too_many_lambdas Ts t =
   416     if is_formula_type (fastype_of1 (Ts, t)) then
   417       exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
   418     else
   419       term_has_too_many_lambdas max_lambda_nesting t
   420 
   421 (* The max apply depth of any "metis" call in "Metis_Examples" (on 31-10-2007)
   422    was 11. *)
   423 val max_apply_depth = 15
   424 
   425 fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
   426   | apply_depth (Abs (_, _, t)) = apply_depth t
   427   | apply_depth _ = 0
   428 
   429 fun is_formula_too_complex t =
   430   apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
   431 
   432 val exists_sledgehammer_const =
   433   exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
   434 
   435 fun is_strange_thm th =
   436   case head_of (concl_of th) of
   437       Const (a, _) => (a <> @{const_name Trueprop} andalso
   438                        a <> @{const_name "=="})
   439     | _ => false
   440 
   441 val type_has_top_sort =
   442   exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
   443 
   444 fun is_theorem_bad_for_atps thm =
   445   let val t = prop_of thm in
   446     is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
   447     exists_sledgehammer_const t orelse is_strange_thm thm
   448   end
   449 
   450 fun all_name_thms_pairs ctxt chained_ths =
   451   let
   452     val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt);
   453     val local_facts = ProofContext.facts_of ctxt;
   454     val full_space =
   455       Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts);
   456 
   457     fun valid_facts facts =
   458       (facts, []) |-> Facts.fold_static (fn (name, ths0) =>
   459         if Facts.is_concealed facts name orelse
   460            (respect_no_atp andalso is_package_def name) orelse
   461            member (op =) multi_base_blacklist (Long_Name.base_name name) orelse
   462            String.isSuffix "_def_raw" (* FIXME: crude hack *) name then
   463           I
   464         else
   465           let
   466             fun check_thms a =
   467               (case try (ProofContext.get_thms ctxt) a of
   468                 NONE => false
   469               | SOME ths1 => Thm.eq_thms (ths0, ths1));
   470 
   471             val name1 = Facts.extern facts name;
   472             val name2 = Name_Space.extern full_space name;
   473             val ths = filter_out is_theorem_bad_for_atps ths0
   474           in
   475             case find_first check_thms [name1, name2, name] of
   476               NONE => I
   477             | SOME name' =>
   478               cons (name' |> forall (member Thm.eq_thm chained_ths) ths
   479                              ? prefix chained_prefix, ths)
   480           end)
   481   in valid_facts global_facts @ valid_facts local_facts end;
   482 
   483 fun multi_name a th (n, pairs) =
   484   (n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs);
   485 
   486 fun add_names (_, []) pairs = pairs
   487   | add_names (a, [th]) pairs = (a, th) :: pairs
   488   | add_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs))
   489 
   490 fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a;
   491 
   492 (* The single-name theorems go after the multiple-name ones, so that single
   493    names are preferred when both are available. *)
   494 fun name_thm_pairs ctxt respect_no_atp name_thms_pairs =
   495   let
   496     val (mults, singles) = List.partition is_multi name_thms_pairs
   497     val ps = [] |> fold add_names singles |> fold add_names mults
   498   in ps |> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd) end;
   499 
   500 fun is_named ("", th) =
   501     (warning ("No name for theorem " ^
   502               Display.string_of_thm_without_context th); false)
   503   | is_named _ = true
   504 fun checked_name_thm_pairs respect_no_atp ctxt =
   505   name_thm_pairs ctxt respect_no_atp
   506   #> tap (fn ps => trace_msg
   507                         (fn () => ("Considering " ^ Int.toString (length ps) ^
   508                                    " theorems")))
   509   #> filter is_named
   510 
   511 (***************************************************************)
   512 (* ATP invocation methods setup                                *)
   513 (***************************************************************)
   514 
   515 (**** Predicates to detect unwanted facts (prolific or likely to cause
   516       unsoundness) ****)
   517 
   518 (* Too general means, positive equality literal with a variable X as one
   519    operand, when X does not occur properly in the other operand. This rules out
   520    clearly inconsistent facts such as X = a | X = b, though it by no means
   521    guarantees soundness. *)
   522 
   523 (* Unwanted equalities are those between a (bound or schematic) variable that
   524    does not properly occur in the second operand. *)
   525 fun too_general_eqterms (Var z) t =
   526     not (exists_subterm (fn Var z' => z = z' | _ => false) t)
   527   | too_general_eqterms (Bound j) t = not (loose_bvar1 (t, j))
   528   | too_general_eqterms _ _ = false
   529 
   530 fun too_general_equality (Const (@{const_name "op ="}, _) $ t1 $ t2) =
   531     too_general_eqterms t1 t2 orelse too_general_eqterms t2 t1
   532   | too_general_equality _ = false
   533 
   534 fun has_bound_or_var_of_type tycons =
   535   exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s
   536                    | Abs (_, Type (s, _), _) => member (op =) tycons s
   537                    | _ => false)
   538 
   539 (* Facts are forbidden to contain variables of these types. The typical reason
   540    is that they lead to unsoundness. Note that "unit" satisfies numerous
   541    equations like "?x = ()". The resulting clauses will have no type constraint,
   542    yielding false proofs. Even "bool" leads to many unsound proofs, though only
   543    for higher-order problems. *)
   544 val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}];
   545 
   546 (* Facts containing variables of type "unit" or "bool" or of the form
   547    "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
   548    are omitted. *)
   549 fun is_dangerous_term full_types t =
   550     not full_types andalso
   551     (has_bound_or_var_of_type dangerous_types t orelse
   552      forall too_general_equality (HOLogic.disjuncts (strip_Trueprop_and_all t)))
   553 
   554 fun relevant_facts full_types relevance_threshold relevance_convergence
   555                    defs_relevant max_new theory_relevant
   556                    (relevance_override as {add, del, only})
   557                    (ctxt, (chained_ths, _)) hyp_ts concl_t =
   558   let
   559     val add_thms = maps (ProofContext.get_fact ctxt) add
   560     val axioms =
   561       checked_name_thm_pairs (respect_no_atp andalso not only) ctxt
   562           (map (name_thms_pair_from_ref ctxt chained_ths) add @
   563            (if only then [] else all_name_thms_pairs ctxt chained_ths))
   564       |> filter (fn (_, th) => member Thm.eq_thm add_thms th orelse
   565                                not (is_dangerous_term full_types (prop_of th)))
   566   in
   567     relevance_filter ctxt relevance_threshold relevance_convergence
   568                      defs_relevant max_new theory_relevant relevance_override
   569                      axioms (concl_t :: hyp_ts)
   570     |> make_unique |> sort_wrt fst
   571   end
   572 
   573 end;