src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
author blanchet
Wed, 25 Aug 2010 19:41:18 +0200
changeset 38984 ad577fd62ee4
parent 38983 2b6333f78a9e
child 38986 b264ae66cede
permissions -rw-r--r--
reorganize options regarding to the relevance threshold and decay
     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 name_thm_pairs_from_ref :
    15     Proof.context -> unit Symtab.table -> thm list -> Facts.ref
    16     -> ((unit -> string * bool) * (bool * thm)) list
    17   val relevant_facts :
    18     bool -> real * real -> int -> bool -> relevance_override
    19     -> Proof.context * (thm list * 'a) -> term list -> term
    20     -> ((string * bool) * thm) list
    21 end;
    22 
    23 structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
    24 struct
    25 
    26 open Sledgehammer_Util
    27 
    28 val trace = Unsynchronized.ref false
    29 fun trace_msg msg = if !trace then tracing (msg ()) else ()
    30 
    31 val respect_no_atp = true
    32 
    33 type relevance_override =
    34   {add: Facts.ref list,
    35    del: Facts.ref list,
    36    only: bool}
    37 
    38 val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
    39 
    40 fun repair_name reserved multi j name =
    41   (name |> Symtab.defined reserved name ? quote) ^
    42   (if multi then "(" ^ Int.toString j ^ ")" else "")
    43 
    44 fun name_thm_pairs_from_ref ctxt reserved chained_ths xref =
    45   let
    46     val ths = ProofContext.get_fact ctxt xref
    47     val name = Facts.string_of_ref xref
    48     val multi = length ths > 1
    49   in
    50     fold (fn th => fn (j, rest) =>
    51              (j + 1, (fn () => (repair_name reserved multi j name,
    52                                 member Thm.eq_thm chained_ths th),
    53                       (multi, th)) :: rest))
    54          ths (1, [])
    55     |> snd
    56   end
    57 
    58 (***************************************************************)
    59 (* Relevance Filtering                                         *)
    60 (***************************************************************)
    61 
    62 (*** constants with types ***)
    63 
    64 (*An abstraction of Isabelle types*)
    65 datatype pseudotype = PVar | PType of string * pseudotype list
    66 
    67 fun string_for_pseudotype PVar = "?"
    68   | string_for_pseudotype (PType (s, Ts)) =
    69     (case Ts of
    70        [] => ""
    71      | [T] => string_for_pseudotype T
    72      | Ts => string_for_pseudotypes Ts ^ " ") ^ s
    73 and string_for_pseudotypes Ts =
    74   "(" ^ commas (map string_for_pseudotype Ts) ^ ")"
    75 
    76 (*Is the second type an instance of the first one?*)
    77 fun match_pseudotype (PType (a, T), PType (b, U)) =
    78     a = b andalso match_pseudotypes (T, U)
    79   | match_pseudotype (PVar, _) = true
    80   | match_pseudotype (_, PVar) = false
    81 and match_pseudotypes ([], []) = true
    82   | match_pseudotypes (T :: Ts, U :: Us) =
    83     match_pseudotype (T, U) andalso match_pseudotypes (Ts, Us)
    84 
    85 (*Is there a unifiable constant?*)
    86 fun pseudoconst_mem f const_tab (c, c_typ) =
    87   exists (curry (match_pseudotypes o f) c_typ)
    88          (these (Symtab.lookup const_tab c))
    89 
    90 fun pseudotype_for (Type (c,typs)) = PType (c, map pseudotype_for typs)
    91   | pseudotype_for (TFree _) = PVar
    92   | pseudotype_for (TVar _) = PVar
    93 (* Pairs a constant with the list of its type instantiations. *)
    94 fun pseudoconst_for thy (c, T) =
    95   (c, map pseudotype_for (Sign.const_typargs thy (c, T)))
    96   handle TYPE _ => (c, [])  (* Variable (locale constant): monomorphic *)
    97 
    98 fun string_for_pseudoconst (s, []) = s
    99   | string_for_pseudoconst (s, Ts) = s ^ string_for_pseudotypes Ts
   100 fun string_for_super_pseudoconst (s, [[]]) = s
   101   | string_for_super_pseudoconst (s, Tss) =
   102     s ^ "{" ^ commas (map string_for_pseudotypes Tss) ^ "}"
   103 
   104 (*Add a const/type pair to the table, but a [] entry means a standard connective,
   105   which we ignore.*)
   106 fun add_const_to_table (c, ctyps) =
   107   Symtab.map_default (c, [ctyps])
   108                      (fn [] => [] | ctypss => insert (op =) ctyps ctypss)
   109 
   110 fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
   111 
   112 val fresh_prefix = "Sledgehammer.skolem."
   113 val flip = Option.map not
   114 (* These are typically simplified away by "Meson.presimplify". *)
   115 val boring_consts =
   116   [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let}]
   117 
   118 fun get_consts thy pos ts =
   119   let
   120     (* We include free variables, as well as constants, to handle locales. For
   121        each quantifiers that must necessarily be skolemized by the ATP, we
   122        introduce a fresh constant to simulate the effect of Skolemization. *)
   123     fun do_term t =
   124       case t of
   125         Const x => add_const_to_table (pseudoconst_for thy x)
   126       | Free (s, _) => add_const_to_table (s, [])
   127       | t1 $ t2 => fold do_term [t1, t2]
   128       | Abs (_, _, t') => do_term t'
   129       | _ => I
   130     fun do_quantifier will_surely_be_skolemized body_t =
   131       do_formula pos body_t
   132       #> (if will_surely_be_skolemized then
   133             add_const_to_table (gensym fresh_prefix, [])
   134           else
   135             I)
   136     and do_term_or_formula T =
   137       if is_formula_type T then do_formula NONE else do_term
   138     and do_formula pos t =
   139       case t of
   140         Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
   141         do_quantifier (pos = SOME false) body_t
   142       | @{const "==>"} $ t1 $ t2 =>
   143         do_formula (flip pos) t1 #> do_formula pos t2
   144       | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
   145         fold (do_term_or_formula T) [t1, t2]
   146       | @{const Trueprop} $ t1 => do_formula pos t1
   147       | @{const Not} $ t1 => do_formula (flip pos) t1
   148       | Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
   149         do_quantifier (pos = SOME false) body_t
   150       | Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
   151         do_quantifier (pos = SOME true) body_t
   152       | @{const "op &"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
   153       | @{const "op |"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
   154       | @{const "op -->"} $ t1 $ t2 =>
   155         do_formula (flip pos) t1 #> do_formula pos t2
   156       | Const (@{const_name "op ="}, Type (_, [T, _])) $ t1 $ t2 =>
   157         fold (do_term_or_formula T) [t1, t2]
   158       | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
   159         $ t1 $ t2 $ t3 =>
   160         do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
   161       | Const (@{const_name Ex1}, _) $ Abs (_, _, body_t) =>
   162         do_quantifier (is_some pos) body_t
   163       | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, body_t) =>
   164         do_quantifier (pos = SOME false)
   165                       (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, body_t))
   166       | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, body_t) =>
   167         do_quantifier (pos = SOME true)
   168                       (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, body_t))
   169       | (t0 as Const (_, @{typ bool})) $ t1 =>
   170         do_term t0 #> do_formula pos t1  (* theory constant *)
   171       | _ => do_term t
   172   in
   173     Symtab.empty |> fold (Symtab.update o rpair []) boring_consts
   174                  |> fold (do_formula pos) ts
   175   end
   176 
   177 (*Inserts a dummy "constant" referring to the theory name, so that relevance
   178   takes the given theory into account.*)
   179 fun theory_const_prop_of theory_relevant th =
   180   if theory_relevant then
   181     let
   182       val name = Context.theory_name (theory_of_thm th)
   183       val t = Const (name ^ ". 1", @{typ bool})
   184     in t $ prop_of th end
   185   else
   186     prop_of th
   187 
   188 (**** Constant / Type Frequencies ****)
   189 
   190 (* A two-dimensional symbol table counts frequencies of constants. It's keyed
   191    first by constant name and second by its list of type instantiations. For the
   192    latter, we need a linear ordering on "pseudotype list". *)
   193 
   194 fun pseudotype_ord p =
   195   case p of
   196     (PVar, PVar) => EQUAL
   197   | (PVar, PType _) => LESS
   198   | (PType _, PVar) => GREATER
   199   | (PType q1, PType q2) =>
   200     prod_ord fast_string_ord (dict_ord pseudotype_ord) (q1, q2)
   201 
   202 structure CTtab =
   203   Table(type key = pseudotype list val ord = dict_ord pseudotype_ord)
   204 
   205 fun count_axiom_consts theory_relevant thy (_, th) =
   206   let
   207     fun do_const (a, T) =
   208       let val (c, cts) = pseudoconst_for thy (a, T) in
   209         (* Two-dimensional table update. Constant maps to types maps to
   210            count. *)
   211         CTtab.map_default (cts, 0) (Integer.add 1)
   212         |> Symtab.map_default (c, CTtab.empty)
   213       end
   214     fun do_term (Const x) = do_const x
   215       | do_term (Free x) = do_const x
   216       | do_term (t $ u) = do_term t #> do_term u
   217       | do_term (Abs (_, _, t)) = do_term t
   218       | do_term _ = I
   219   in th |> theory_const_prop_of theory_relevant |> do_term end
   220 
   221 
   222 (**** Actual Filtering Code ****)
   223 
   224 (*The frequency of a constant is the sum of those of all instances of its type.*)
   225 fun pseudoconst_freq match const_tab (c, cts) =
   226   CTtab.fold (fn (cts', m) => match (cts, cts') ? Integer.add m)
   227              (the (Symtab.lookup const_tab c)) 0
   228   handle Option.Option => 0
   229 
   230 
   231 (* A surprising number of theorems contain only a few significant constants.
   232    These include all induction rules, and other general theorems. *)
   233 
   234 (* "log" seems best in practice. A constant function of one ignores the constant
   235    frequencies. *)
   236 fun rel_log (x : real) = 1.0 + 2.0 / Math.ln (x + 1.0)
   237 fun irrel_log (x : real) = Math.ln (x + 19.0) / 6.4
   238 
   239 (* Computes a constant's weight, as determined by its frequency. *)
   240 val rel_weight = rel_log o real oo pseudoconst_freq match_pseudotypes
   241 val irrel_weight =
   242   irrel_log o real oo pseudoconst_freq (match_pseudotypes o swap)
   243 (* fun irrel_weight _ _ = 1.0  FIXME: OLD CODE *)
   244 
   245 fun axiom_weight const_tab relevant_consts axiom_consts =
   246   case axiom_consts |> List.partition (pseudoconst_mem I relevant_consts)
   247                     ||> filter_out (pseudoconst_mem swap relevant_consts) of
   248     ([], []) => 0.0
   249   | (_, []) => 1.0
   250   | (rel, irrel) =>
   251     let
   252       val rel_weight = fold (curry Real.+ o rel_weight const_tab) rel 0.0
   253       val irrel_weight = fold (curry Real.+ o irrel_weight const_tab) irrel 0.0
   254       val res = rel_weight / (rel_weight + irrel_weight)
   255     in if Real.isFinite res then res else 0.0 end
   256 
   257 fun consts_of_term thy t =
   258   Symtab.fold (fn (x, ys) => fold (fn y => cons (x, y)) ys)
   259               (get_consts thy (SOME true) [t]) []
   260 
   261 fun pair_consts_axiom theory_relevant thy axiom =
   262   (axiom, axiom |> snd |> theory_const_prop_of theory_relevant
   263                 |> consts_of_term thy)
   264 
   265 type annotated_thm =
   266   ((unit -> string * bool) * thm) * (string * pseudotype list) list
   267 
   268 fun take_best max (candidates : (annotated_thm * real) list) =
   269   let
   270     val ((perfect, more_perfect), imperfect) =
   271       candidates |> List.partition (fn (_, w) => w > 0.99999) |>> chop (max - 1)
   272                  ||> sort (Real.compare o swap o pairself snd)
   273     val (accepts, rejects) =
   274       case more_perfect @ imperfect of
   275         [] => (perfect, [])
   276       | (q :: qs) => (q :: perfect, qs)
   277   in
   278     trace_msg (fn () => "Number of candidates: " ^
   279                         string_of_int (length candidates));
   280     trace_msg (fn () => "Effective threshold: " ^
   281                         Real.toString (#2 (hd accepts)));
   282     trace_msg (fn () => "Actually passed: " ^
   283         (accepts |> map (fn (((name, _), _), weight) =>
   284                             fst (name ()) ^ " [" ^ Real.toString weight ^ "]")
   285                  |> commas));
   286     (accepts, rejects)
   287   end
   288 
   289 val threshold_divisor = 2.0
   290 val ridiculous_threshold = 0.1
   291 
   292 fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant
   293                      ({add, del, ...} : relevance_override) axioms goal_ts =
   294   let
   295     val thy = ProofContext.theory_of ctxt
   296     val const_tab = fold (count_axiom_consts theory_relevant thy) axioms
   297                          Symtab.empty
   298     val add_thms = maps (ProofContext.get_fact ctxt) add
   299     val del_thms = maps (ProofContext.get_fact ctxt) del
   300     fun iter j max threshold rel_const_tab hopeless hopeful =
   301       let
   302         fun game_over rejects =
   303           if j = 0 andalso threshold >= ridiculous_threshold then
   304             (* First iteration? Try again. *)
   305             iter 0 max (threshold / threshold_divisor) rel_const_tab hopeless
   306                  hopeful
   307           else
   308             (* Add "add:" facts. *)
   309             if null add_thms then
   310               []
   311             else
   312               map_filter (fn ((p as (_, th), _), _) =>
   313                              if member Thm.eq_thm add_thms th then SOME p
   314                              else NONE) rejects
   315         fun relevant [] rejects [] hopeless =
   316             (* Nothing has been added this iteration. *)
   317             game_over (map (apsnd SOME) (rejects @ hopeless))
   318           | relevant candidates rejects [] hopeless =
   319             let
   320               val (accepts, more_rejects) = take_best max candidates
   321               val rel_const_tab' =
   322                 rel_const_tab
   323                 |> fold add_const_to_table (maps (snd o fst) accepts)
   324               fun is_dirty (c, _) =
   325                 Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
   326               val (hopeful_rejects, hopeless_rejects) =
   327                  (rejects @ hopeless, ([], []))
   328                  |-> fold (fn (ax as (_, consts), old_weight) =>
   329                               if exists is_dirty consts then
   330                                 apfst (cons (ax, NONE))
   331                               else
   332                                 apsnd (cons (ax, old_weight)))
   333                  |>> append (more_rejects
   334                              |> map (fn (ax as (_, consts), old_weight) =>
   335                                         (ax, if exists is_dirty consts then NONE
   336                                              else SOME old_weight)))
   337               val threshold = threshold + (1.0 - threshold) * decay
   338               val max = max - length accepts
   339             in
   340               trace_msg (fn () => "New or updated constants: " ^
   341                   commas (rel_const_tab' |> Symtab.dest
   342                           |> subtract (op =) (Symtab.dest rel_const_tab)
   343                           |> map string_for_super_pseudoconst));
   344               map (fst o fst) accepts @
   345               (if max = 0 then
   346                  game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
   347                else
   348                  iter (j + 1) max threshold rel_const_tab' hopeless_rejects
   349                       hopeful_rejects)
   350             end
   351           | relevant candidates rejects
   352                      (((ax as ((name, th), axiom_consts)), cached_weight)
   353                       :: hopeful) hopeless =
   354             let
   355               val weight =
   356                 case cached_weight of
   357                   SOME w => w
   358                 | NONE => axiom_weight const_tab rel_const_tab axiom_consts
   359             in
   360               if weight >= threshold then
   361                 relevant ((ax, weight) :: candidates) rejects hopeful hopeless
   362               else
   363                 relevant candidates ((ax, weight) :: rejects) hopeful hopeless
   364             end
   365         in
   366           trace_msg (fn () =>
   367               "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
   368               Real.toString threshold ^ ", constants: " ^
   369               commas (rel_const_tab |> Symtab.dest
   370                       |> filter (curry (op <>) [] o snd)
   371                       |> map string_for_super_pseudoconst));
   372           relevant [] [] hopeful hopeless
   373         end
   374   in
   375     axioms |> filter_out (member Thm.eq_thm del_thms o snd)
   376            |> map (rpair NONE o pair_consts_axiom theory_relevant thy)
   377            |> iter 0 max_relevant threshold0
   378                    (get_consts thy (SOME false) goal_ts) []
   379            |> tap (fn res => trace_msg (fn () =>
   380                                 "Total relevant: " ^ Int.toString (length res)))
   381   end
   382 
   383 
   384 (***************************************************************)
   385 (* Retrieving and filtering lemmas                             *)
   386 (***************************************************************)
   387 
   388 (*** retrieve lemmas and filter them ***)
   389 
   390 (*Reject theorems with names like "List.filter.filter_list_def" or
   391   "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
   392 fun is_package_def a =
   393   let val names = Long_Name.explode a
   394   in
   395      length names > 2 andalso
   396      not (hd names = "local") andalso
   397      String.isSuffix "_def" a  orelse  String.isSuffix "_defs" a
   398   end;
   399 
   400 fun make_fact_table xs =
   401   fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty
   402 fun make_unique xs = Termtab.fold (cons o snd) (make_fact_table xs) []
   403 
   404 (* FIXME: put other record thms here, or declare as "no_atp" *)
   405 val multi_base_blacklist =
   406   ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
   407    "split_asm", "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy",
   408    "case_cong", "weak_case_cong"]
   409   |> map (prefix ".")
   410 
   411 val max_lambda_nesting = 3
   412 
   413 fun term_has_too_many_lambdas max (t1 $ t2) =
   414     exists (term_has_too_many_lambdas max) [t1, t2]
   415   | term_has_too_many_lambdas max (Abs (_, _, t)) =
   416     max = 0 orelse term_has_too_many_lambdas (max - 1) t
   417   | term_has_too_many_lambdas _ _ = false
   418 
   419 (* Don't count nested lambdas at the level of formulas, since they are
   420    quantifiers. *)
   421 fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
   422     formula_has_too_many_lambdas (T :: Ts) t
   423   | formula_has_too_many_lambdas Ts t =
   424     if is_formula_type (fastype_of1 (Ts, t)) then
   425       exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
   426     else
   427       term_has_too_many_lambdas max_lambda_nesting t
   428 
   429 (* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31)
   430    was 11. *)
   431 val max_apply_depth = 15
   432 
   433 fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
   434   | apply_depth (Abs (_, _, t)) = apply_depth t
   435   | apply_depth _ = 0
   436 
   437 fun is_formula_too_complex t =
   438   apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
   439 
   440 val exists_sledgehammer_const =
   441   exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
   442 
   443 fun is_strange_theorem th =
   444   case head_of (concl_of th) of
   445       Const (a, _) => (a <> @{const_name Trueprop} andalso
   446                        a <> @{const_name "=="})
   447     | _ => false
   448 
   449 val type_has_top_sort =
   450   exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
   451 
   452 (**** Predicates to detect unwanted facts (prolific or likely to cause
   453       unsoundness) ****)
   454 
   455 (* Too general means, positive equality literal with a variable X as one
   456    operand, when X does not occur properly in the other operand. This rules out
   457    clearly inconsistent facts such as X = a | X = b, though it by no means
   458    guarantees soundness. *)
   459 
   460 (* Unwanted equalities are those between a (bound or schematic) variable that
   461    does not properly occur in the second operand. *)
   462 val is_exhaustive_finite =
   463   let
   464     fun is_bad_equal (Var z) t =
   465         not (exists_subterm (fn Var z' => z = z' | _ => false) t)
   466       | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
   467       | is_bad_equal _ _ = false
   468     fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
   469     fun do_formula pos t =
   470       case (pos, t) of
   471         (_, @{const Trueprop} $ t1) => do_formula pos t1
   472       | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
   473         do_formula pos t'
   474       | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
   475         do_formula pos t'
   476       | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
   477         do_formula pos t'
   478       | (_, @{const "==>"} $ t1 $ t2) =>
   479         do_formula (not pos) t1 andalso
   480         (t2 = @{prop False} orelse do_formula pos t2)
   481       | (_, @{const "op -->"} $ t1 $ t2) =>
   482         do_formula (not pos) t1 andalso
   483         (t2 = @{const False} orelse do_formula pos t2)
   484       | (_, @{const Not} $ t1) => do_formula (not pos) t1
   485       | (true, @{const "op |"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
   486       | (false, @{const "op &"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
   487       | (true, Const (@{const_name "op ="}, _) $ t1 $ t2) => do_equals t1 t2
   488       | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
   489       | _ => false
   490   in do_formula true end
   491 
   492 fun has_bound_or_var_of_type tycons =
   493   exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s
   494                    | Abs (_, Type (s, _), _) => member (op =) tycons s
   495                    | _ => false)
   496 
   497 (* Facts are forbidden to contain variables of these types. The typical reason
   498    is that they lead to unsoundness. Note that "unit" satisfies numerous
   499    equations like "?x = ()". The resulting clauses will have no type constraint,
   500    yielding false proofs. Even "bool" leads to many unsound proofs, though only
   501    for higher-order problems. *)
   502 val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}];
   503 
   504 (* Facts containing variables of type "unit" or "bool" or of the form
   505    "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
   506    are omitted. *)
   507 fun is_dangerous_term full_types t =
   508   not full_types andalso
   509   let val t = transform_elim_term t in
   510     has_bound_or_var_of_type dangerous_types t orelse
   511     is_exhaustive_finite t
   512   end
   513 
   514 fun is_theorem_bad_for_atps full_types thm =
   515   let val t = prop_of thm in
   516     is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
   517     is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
   518     is_strange_theorem thm
   519   end
   520 
   521 fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths =
   522   let
   523     val is_chained = member Thm.eq_thm chained_ths
   524     val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt)
   525     val local_facts = ProofContext.facts_of ctxt
   526     val named_locals = local_facts |> Facts.dest_static []
   527     (* Unnamed, not chained formulas with schematic variables are omitted,
   528        because they are rejected by the backticks (`...`) parser for some
   529        reason. *)
   530     fun is_good_unnamed_local th =
   531       forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals
   532       andalso (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th))
   533     val unnamed_locals =
   534       local_facts |> Facts.props |> filter is_good_unnamed_local
   535                   |> map (pair "" o single)
   536     val full_space =
   537       Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
   538     fun add_valid_facts foldx facts =
   539       foldx (fn (name0, ths) =>
   540         if name0 <> "" andalso
   541            forall (not o member Thm.eq_thm add_thms) ths andalso
   542            (Facts.is_concealed facts name0 orelse
   543             (respect_no_atp andalso is_package_def name0) orelse
   544             exists (fn s => String.isSuffix s name0) multi_base_blacklist orelse
   545             String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then
   546           I
   547         else
   548           let
   549             val multi = length ths > 1
   550             fun backquotify th =
   551               "`" ^ Print_Mode.setmp [Print_Mode.input]
   552                                  (Syntax.string_of_term ctxt) (prop_of th) ^ "`"
   553               |> String.translate (fn c => if Char.isPrint c then str c else "")
   554               |> simplify_spaces
   555             fun check_thms a =
   556               case try (ProofContext.get_thms ctxt) a of
   557                 NONE => false
   558               | SOME ths' => Thm.eq_thms (ths, ths')
   559           in
   560             pair 1
   561             #> fold (fn th => fn (j, rest) =>
   562                  (j + 1,
   563                   if is_theorem_bad_for_atps full_types th andalso
   564                      not (member Thm.eq_thm add_thms th) then
   565                     rest
   566                   else
   567                     (fn () =>
   568                         (if name0 = "" then
   569                            th |> backquotify
   570                          else
   571                            let
   572                              val name1 = Facts.extern facts name0
   573                              val name2 = Name_Space.extern full_space name0
   574                            in
   575                              case find_first check_thms [name1, name2, name0] of
   576                                SOME name => repair_name reserved multi j name
   577                              | NONE => ""
   578                            end, is_chained th), (multi, th)) :: rest)) ths
   579             #> snd
   580           end)
   581   in
   582     [] |> add_valid_facts fold local_facts (unnamed_locals @ named_locals)
   583        |> add_valid_facts Facts.fold_static global_facts global_facts
   584   end
   585 
   586 (* The single-name theorems go after the multiple-name ones, so that single
   587    names are preferred when both are available. *)
   588 fun name_thm_pairs ctxt respect_no_atp =
   589   List.partition (fst o snd) #> op @ #> map (apsnd snd)
   590   #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd)
   591 
   592 (***************************************************************)
   593 (* ATP invocation methods setup                                *)
   594 (***************************************************************)
   595 
   596 fun relevant_facts full_types (threshold0, threshold1) max_relevant
   597                    theory_relevant (relevance_override as {add, del, only})
   598                    (ctxt, (chained_ths, _)) hyp_ts concl_t =
   599   let
   600     val decay = 1.0 - Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
   601                                 1.0 / Real.fromInt (max_relevant + 1))
   602     val add_thms = maps (ProofContext.get_fact ctxt) add
   603     val reserved = reserved_isar_keyword_table ()
   604     val axioms =
   605       (if only then
   606          maps (name_thm_pairs_from_ref ctxt reserved chained_ths) add
   607        else
   608          all_name_thms_pairs ctxt reserved full_types add_thms chained_ths)
   609       |> name_thm_pairs ctxt (respect_no_atp andalso not only)
   610       |> make_unique
   611   in
   612     trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^
   613                         " theorems");
   614     (if threshold0 > 1.0 orelse threshold0 > threshold1 then
   615        []
   616      else if threshold0 < 0.0 then
   617        axioms
   618      else
   619        relevance_filter ctxt threshold0 decay max_relevant theory_relevant
   620                         relevance_override axioms (concl_t :: hyp_ts))
   621     |> map (apfst (fn f => f ())) |> sort_wrt (fst o fst)
   622   end
   623 
   624 end;