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