src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML
author blanchet
Fri, 20 Jul 2012 22:19:45 +0200
changeset 49396 1b7d798460bb
parent 49395 d4b7c7be3116
child 49421 b002cc16aa99
permissions -rw-r--r--
renamed ML structures
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_mepo.ML
     2     Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
     3     Author:     Jasmin Blanchette, TU Muenchen
     4 
     5 Sledgehammer's iterative relevance filter (MePo = Meng-Paulson).
     6 *)
     7 
     8 signature SLEDGEHAMMER_MEPO =
     9 sig
    10   type stature = ATP_Problem_Generate.stature
    11   type fact = Sledgehammer_Fact.fact
    12   type params = Sledgehammer_Provers.params
    13   type relevance_fudge = Sledgehammer_Provers.relevance_fudge
    14 
    15   val trace : bool Config.T
    16   val pseudo_abs_name : string
    17   val pseudo_skolem_prefix : string
    18   val const_names_in_fact :
    19     theory -> (string * typ -> term list -> bool * term list) -> term
    20     -> string list
    21   val iterative_relevant_facts :
    22     Proof.context -> params -> string -> int -> relevance_fudge option
    23     -> term list -> term -> fact list -> fact list
    24 end;
    25 
    26 structure Sledgehammer_MePo : SLEDGEHAMMER_MEPO =
    27 struct
    28 
    29 open ATP_Problem_Generate
    30 open Sledgehammer_Fact
    31 open Sledgehammer_Provers
    32 
    33 val trace =
    34   Attrib.setup_config_bool @{binding sledgehammer_filter_iter_trace} (K false)
    35 fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
    36 
    37 val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
    38 val pseudo_abs_name = sledgehammer_prefix ^ "abs"
    39 val pseudo_skolem_prefix = sledgehammer_prefix ^ "sko"
    40 val theory_const_suffix = Long_Name.separator ^ " 1"
    41 
    42 fun order_of_type (Type (@{type_name fun}, [T1, T2])) =
    43     Int.max (order_of_type T1 + 1, order_of_type T2)
    44   | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
    45   | order_of_type _ = 0
    46 
    47 (* An abstraction of Isabelle types and first-order terms *)
    48 datatype pattern = PVar | PApp of string * pattern list
    49 datatype ptype = PType of int * pattern list
    50 
    51 fun string_for_pattern PVar = "_"
    52   | string_for_pattern (PApp (s, ps)) =
    53     if null ps then s else s ^ string_for_patterns ps
    54 and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
    55 fun string_for_ptype (PType (_, ps)) = string_for_patterns ps
    56 
    57 (*Is the second type an instance of the first one?*)
    58 fun match_pattern (PVar, _) = true
    59   | match_pattern (PApp _, PVar) = false
    60   | match_pattern (PApp (s, ps), PApp (t, qs)) =
    61     s = t andalso match_patterns (ps, qs)
    62 and match_patterns (_, []) = true
    63   | match_patterns ([], _) = false
    64   | match_patterns (p :: ps, q :: qs) =
    65     match_pattern (p, q) andalso match_patterns (ps, qs)
    66 fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
    67 
    68 (* Is there a unifiable constant? *)
    69 fun pconst_mem f consts (s, ps) =
    70   exists (curry (match_ptype o f) ps)
    71          (map snd (filter (curry (op =) s o fst) consts))
    72 fun pconst_hyper_mem f const_tab (s, ps) =
    73   exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
    74 
    75 fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts)
    76   | pattern_for_type (TFree (s, _)) = PApp (s, [])
    77   | pattern_for_type (TVar _) = PVar
    78 
    79 (* Pairs a constant with the list of its type instantiations. *)
    80 fun ptype thy const x =
    81   (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x))
    82    else [])
    83 fun rich_ptype thy const (s, T) =
    84   PType (order_of_type T, ptype thy const (s, T))
    85 fun rich_pconst thy const (s, T) = (s, rich_ptype thy const (s, T))
    86 
    87 fun string_for_hyper_pconst (s, ps) =
    88   s ^ "{" ^ commas (map string_for_ptype ps) ^ "}"
    89 
    90 (* Add a pconstant to the table, but a [] entry means a standard
    91    connective, which we ignore.*)
    92 fun add_pconst_to_table also_skolem (s, p) =
    93   if (not also_skolem andalso String.isPrefix pseudo_skolem_prefix s) then I
    94   else Symtab.map_default (s, [p]) (insert (op =) p)
    95 
    96 (* Set constants tend to pull in too many irrelevant facts. We limit the damage
    97    by treating them more or less as if they were built-in but add their
    98    axiomatization at the end. *)
    99 val set_consts = [@{const_name Collect}, @{const_name Set.member}]
   100 val set_thms = @{thms Collect_mem_eq mem_Collect_eq Collect_cong}
   101 
   102 fun add_pconsts_in_term thy is_built_in_const also_skolems pos =
   103   let
   104     val flip = Option.map not
   105     (* We include free variables, as well as constants, to handle locales. For
   106        each quantifiers that must necessarily be skolemized by the automatic
   107        prover, we introduce a fresh constant to simulate the effect of
   108        Skolemization. *)
   109     fun do_const const ext_arg (x as (s, _)) ts =
   110       let val (built_in, ts) = is_built_in_const x ts in
   111         if member (op =) set_consts s then
   112           fold (do_term ext_arg) ts
   113         else
   114           (not built_in
   115            ? add_pconst_to_table also_skolems (rich_pconst thy const x))
   116           #> fold (do_term false) ts
   117       end
   118     and do_term ext_arg t =
   119       case strip_comb t of
   120         (Const x, ts) => do_const true ext_arg x ts
   121       | (Free x, ts) => do_const false ext_arg x ts
   122       | (Abs (_, T, t'), ts) =>
   123         ((null ts andalso not ext_arg)
   124          (* Since lambdas on the right-hand side of equalities are usually
   125             extensionalized later by "abs_extensionalize_term", we don't
   126             penalize them here. *)
   127          ? add_pconst_to_table true (pseudo_abs_name,
   128                                      PType (order_of_type T + 1, [])))
   129         #> fold (do_term false) (t' :: ts)
   130       | (_, ts) => fold (do_term false) ts
   131     fun do_quantifier will_surely_be_skolemized abs_T body_t =
   132       do_formula pos body_t
   133       #> (if also_skolems andalso will_surely_be_skolemized then
   134             add_pconst_to_table true (pseudo_skolem_prefix ^ serial_string (),
   135                                       PType (order_of_type abs_T, []))
   136           else
   137             I)
   138     and do_term_or_formula ext_arg T =
   139       if T = HOLogic.boolT then do_formula NONE else do_term ext_arg
   140     and do_formula pos t =
   141       case t of
   142         Const (@{const_name all}, _) $ Abs (_, T, t') =>
   143         do_quantifier (pos = SOME false) T t'
   144       | @{const "==>"} $ t1 $ t2 =>
   145         do_formula (flip pos) t1 #> do_formula pos t2
   146       | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
   147         do_term_or_formula false T t1 #> do_term_or_formula true T t2
   148       | @{const Trueprop} $ t1 => do_formula pos t1
   149       | @{const False} => I
   150       | @{const True} => I
   151       | @{const Not} $ t1 => do_formula (flip pos) t1
   152       | Const (@{const_name All}, _) $ Abs (_, T, t') =>
   153         do_quantifier (pos = SOME false) T t'
   154       | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
   155         do_quantifier (pos = SOME true) T t'
   156       | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
   157       | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
   158       | @{const HOL.implies} $ t1 $ t2 =>
   159         do_formula (flip pos) t1 #> do_formula pos t2
   160       | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
   161         do_term_or_formula false T t1 #> do_term_or_formula true T t2
   162       | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
   163         $ t1 $ t2 $ t3 =>
   164         do_formula NONE t1 #> fold (do_term_or_formula false T) [t2, t3]
   165       | Const (@{const_name Ex1}, _) $ Abs (_, T, t') =>
   166         do_quantifier (is_some pos) T t'
   167       | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
   168         do_quantifier (pos = SOME false) T
   169                       (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t'))
   170       | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
   171         do_quantifier (pos = SOME true) T
   172                       (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t'))
   173       | (t0 as Const (_, @{typ bool})) $ t1 =>
   174         do_term false t0 #> do_formula pos t1  (* theory constant *)
   175       | _ => do_term false t
   176   in do_formula pos end
   177 
   178 fun pconsts_in_fact thy is_built_in_const t =
   179   Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
   180               (Symtab.empty |> add_pconsts_in_term thy is_built_in_const true
   181                                                    (SOME true) t) []
   182 
   183 val const_names_in_fact = map fst ooo pconsts_in_fact
   184 
   185 (* Inserts a dummy "constant" referring to the theory name, so that relevance
   186    takes the given theory into account. *)
   187 fun theory_constify ({theory_const_rel_weight, theory_const_irrel_weight, ...}
   188                      : relevance_fudge) thy_name t =
   189   if exists (curry (op <) 0.0) [theory_const_rel_weight,
   190                                 theory_const_irrel_weight] then
   191     Const (thy_name ^ theory_const_suffix, @{typ bool}) $ t
   192   else
   193     t
   194 
   195 fun theory_const_prop_of fudge th =
   196   theory_constify fudge (Context.theory_name (theory_of_thm th)) (prop_of th)
   197 
   198 fun pair_consts_fact thy is_built_in_const fudge fact =
   199   case fact |> snd |> theory_const_prop_of fudge
   200             |> pconsts_in_fact thy is_built_in_const of
   201     [] => NONE
   202   | consts => SOME ((fact, consts), NONE)
   203 
   204 (* A two-dimensional symbol table counts frequencies of constants. It's keyed
   205    first by constant name and second by its list of type instantiations. For the
   206    latter, we need a linear ordering on "pattern list". *)
   207 
   208 fun pattern_ord p =
   209   case p of
   210     (PVar, PVar) => EQUAL
   211   | (PVar, PApp _) => LESS
   212   | (PApp _, PVar) => GREATER
   213   | (PApp q1, PApp q2) =>
   214     prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
   215 fun ptype_ord (PType p, PType q) =
   216   prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
   217 
   218 structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
   219 
   220 fun count_fact_consts thy fudge =
   221   let
   222     fun do_const const (s, T) ts =
   223       (* Two-dimensional table update. Constant maps to types maps to count. *)
   224       PType_Tab.map_default (rich_ptype thy const (s, T), 0) (Integer.add 1)
   225       |> Symtab.map_default (s, PType_Tab.empty)
   226       #> fold do_term ts
   227     and do_term t =
   228       case strip_comb t of
   229         (Const x, ts) => do_const true x ts
   230       | (Free x, ts) => do_const false x ts
   231       | (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
   232       | (_, ts) => fold do_term ts
   233   in do_term o theory_const_prop_of fudge o snd end
   234 
   235 fun pow_int _ 0 = 1.0
   236   | pow_int x 1 = x
   237   | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x
   238 
   239 (*The frequency of a constant is the sum of those of all instances of its type.*)
   240 fun pconst_freq match const_tab (c, ps) =
   241   PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
   242                  (the (Symtab.lookup const_tab c)) 0
   243 
   244 
   245 (* A surprising number of theorems contain only a few significant constants.
   246    These include all induction rules, and other general theorems. *)
   247 
   248 (* "log" seems best in practice. A constant function of one ignores the constant
   249    frequencies. Rare constants give more points if they are relevant than less
   250    rare ones. *)
   251 fun rel_weight_for _ freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0)
   252 
   253 (* Irrelevant constants are treated differently. We associate lower penalties to
   254    very rare constants and very common ones -- the former because they can't
   255    lead to the inclusion of too many new facts, and the latter because they are
   256    so common as to be of little interest. *)
   257 fun irrel_weight_for ({worse_irrel_freq, higher_order_irrel_weight, ...}
   258                       : relevance_fudge) order freq =
   259   let val (k, x) = worse_irrel_freq |> `Real.ceil in
   260     (if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x
   261      else rel_weight_for order freq / rel_weight_for order k)
   262     * pow_int higher_order_irrel_weight (order - 1)
   263   end
   264 
   265 fun multiplier_for_const_name local_const_multiplier s =
   266   if String.isSubstring "." s then 1.0 else local_const_multiplier
   267 
   268 (* Computes a constant's weight, as determined by its frequency. *)
   269 fun generic_pconst_weight local_const_multiplier abs_weight skolem_weight
   270                           theory_const_weight chained_const_weight weight_for f
   271                           const_tab chained_const_tab (c as (s, PType (m, _))) =
   272   if s = pseudo_abs_name then
   273     abs_weight
   274   else if String.isPrefix pseudo_skolem_prefix s then
   275     skolem_weight
   276   else if String.isSuffix theory_const_suffix s then
   277     theory_const_weight
   278   else
   279     multiplier_for_const_name local_const_multiplier s
   280     * weight_for m (pconst_freq (match_ptype o f) const_tab c)
   281     |> (if chained_const_weight < 1.0 andalso
   282            pconst_hyper_mem I chained_const_tab c then
   283           curry (op *) chained_const_weight
   284         else
   285           I)
   286 
   287 fun rel_pconst_weight ({local_const_multiplier, abs_rel_weight,
   288                         theory_const_rel_weight, ...} : relevance_fudge)
   289                       const_tab =
   290   generic_pconst_weight local_const_multiplier abs_rel_weight 0.0
   291                         theory_const_rel_weight 0.0 rel_weight_for I const_tab
   292                         Symtab.empty
   293 
   294 fun irrel_pconst_weight (fudge as {local_const_multiplier, abs_irrel_weight,
   295                                    skolem_irrel_weight,
   296                                    theory_const_irrel_weight,
   297                                    chained_const_irrel_weight, ...})
   298                         const_tab chained_const_tab =
   299   generic_pconst_weight local_const_multiplier abs_irrel_weight
   300                         skolem_irrel_weight theory_const_irrel_weight
   301                         chained_const_irrel_weight (irrel_weight_for fudge) swap
   302                         const_tab chained_const_tab
   303 
   304 fun stature_bonus ({intro_bonus, ...} : relevance_fudge) (_, Intro) =
   305     intro_bonus
   306   | stature_bonus {elim_bonus, ...} (_, Elim) = elim_bonus
   307   | stature_bonus {simp_bonus, ...} (_, Simp) = simp_bonus
   308   | stature_bonus {local_bonus, ...} (Local, _) = local_bonus
   309   | stature_bonus {assum_bonus, ...} (Assum, _) = assum_bonus
   310   | stature_bonus {chained_bonus, ...} (Chained, _) = chained_bonus
   311   | stature_bonus _ _ = 0.0
   312 
   313 fun is_odd_const_name s =
   314   s = pseudo_abs_name orelse String.isPrefix pseudo_skolem_prefix s orelse
   315   String.isSuffix theory_const_suffix s
   316 
   317 fun fact_weight fudge stature const_tab relevant_consts chained_consts
   318                 fact_consts =
   319   case fact_consts |> List.partition (pconst_hyper_mem I relevant_consts)
   320                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
   321     ([], _) => 0.0
   322   | (rel, irrel) =>
   323     if forall (forall (is_odd_const_name o fst)) [rel, irrel] then
   324       0.0
   325     else
   326       let
   327         val irrel = irrel |> filter_out (pconst_mem swap rel)
   328         val rel_weight =
   329           0.0 |> fold (curry (op +) o rel_pconst_weight fudge const_tab) rel
   330         val irrel_weight =
   331           ~ (stature_bonus fudge stature)
   332           |> fold (curry (op +)
   333                    o irrel_pconst_weight fudge const_tab chained_consts) irrel
   334         val res = rel_weight / (rel_weight + irrel_weight)
   335       in if Real.isFinite res then res else 0.0 end
   336 
   337 fun take_most_relevant ctxt max_facts remaining_max
   338         ({max_imperfect, max_imperfect_exp, ...} : relevance_fudge)
   339         (candidates : ((fact * (string * ptype) list) * real) list) =
   340   let
   341     val max_imperfect =
   342       Real.ceil (Math.pow (max_imperfect,
   343                     Math.pow (Real.fromInt remaining_max
   344                               / Real.fromInt max_facts, max_imperfect_exp)))
   345     val (perfect, imperfect) =
   346       candidates |> sort (Real.compare o swap o pairself snd)
   347                  |> take_prefix (fn (_, w) => w > 0.99999)
   348     val ((accepts, more_rejects), rejects) =
   349       chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
   350   in
   351     trace_msg ctxt (fn () =>
   352         "Actually passed (" ^ string_of_int (length accepts) ^ " of " ^
   353         string_of_int (length candidates) ^ "): " ^
   354         (accepts |> map (fn ((((name, _), _), _), weight) =>
   355                             name () ^ " [" ^ Real.toString weight ^ "]")
   356                  |> commas));
   357     (accepts, more_rejects @ rejects)
   358   end
   359 
   360 fun if_empty_replace_with_scope thy is_built_in_const facts sc tab =
   361   if Symtab.is_empty tab then
   362     Symtab.empty
   363     |> fold (add_pconsts_in_term thy is_built_in_const false (SOME false))
   364             (map_filter (fn ((_, (sc', _)), th) =>
   365                             if sc' = sc then SOME (prop_of th) else NONE) facts)
   366   else
   367     tab
   368 
   369 fun consider_arities is_built_in_const th =
   370   let
   371     fun aux _ _ NONE = NONE
   372       | aux t args (SOME tab) =
   373         case t of
   374           t1 $ t2 => SOME tab |> aux t1 (t2 :: args) |> aux t2 []
   375         | Const (x as (s, _)) =>
   376           (if is_built_in_const x args |> fst then
   377              SOME tab
   378            else case Symtab.lookup tab s of
   379              NONE => SOME (Symtab.update (s, length args) tab)
   380            | SOME n => if n = length args then SOME tab else NONE)
   381         | _ => SOME tab
   382   in aux (prop_of th) [] end
   383 
   384 (* FIXME: This is currently only useful for polymorphic type encodings. *)
   385 fun could_benefit_from_ext is_built_in_const facts =
   386   fold (consider_arities is_built_in_const o snd) facts (SOME Symtab.empty)
   387   |> is_none
   388 
   389 (* High enough so that it isn't wrongly considered as very relevant (e.g., for E
   390    weights), but low enough so that it is unlikely to be truncated away if few
   391    facts are included. *)
   392 val special_fact_index = 75
   393 
   394 fun relevance_filter ctxt thres0 decay max_facts is_built_in_const
   395         (fudge as {threshold_divisor, ridiculous_threshold, ...}) facts hyp_ts
   396         concl_t =
   397   let
   398     val thy = Proof_Context.theory_of ctxt
   399     val const_tab = fold (count_fact_consts thy fudge) facts Symtab.empty
   400     val add_pconsts = add_pconsts_in_term thy is_built_in_const false o SOME
   401     val chained_ts =
   402       facts |> map_filter (fn ((_, (Chained, _)), th) => SOME (prop_of th)
   403                             | _ => NONE)
   404     val chained_const_tab = Symtab.empty |> fold (add_pconsts true) chained_ts
   405     val goal_const_tab =
   406       Symtab.empty |> fold (add_pconsts true) hyp_ts
   407                    |> add_pconsts false concl_t
   408       |> (fn tab => if Symtab.is_empty tab then chained_const_tab else tab)
   409       |> fold (if_empty_replace_with_scope thy is_built_in_const facts)
   410               [Chained, Assum, Local]
   411     fun iter j remaining_max thres rel_const_tab hopeless hopeful =
   412       let
   413         fun relevant [] _ [] =
   414             (* Nothing has been added this iteration. *)
   415             if j = 0 andalso thres >= ridiculous_threshold then
   416               (* First iteration? Try again. *)
   417               iter 0 max_facts (thres / threshold_divisor) rel_const_tab
   418                    hopeless hopeful
   419             else
   420               []
   421           | relevant candidates rejects [] =
   422             let
   423               val (accepts, more_rejects) =
   424                 take_most_relevant ctxt max_facts remaining_max fudge candidates
   425               val rel_const_tab' =
   426                 rel_const_tab
   427                 |> fold (add_pconst_to_table false) (maps (snd o fst) accepts)
   428               fun is_dirty (c, _) =
   429                 Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
   430               val (hopeful_rejects, hopeless_rejects) =
   431                  (rejects @ hopeless, ([], []))
   432                  |-> fold (fn (ax as (_, consts), old_weight) =>
   433                               if exists is_dirty consts then
   434                                 apfst (cons (ax, NONE))
   435                               else
   436                                 apsnd (cons (ax, old_weight)))
   437                  |>> append (more_rejects
   438                              |> map (fn (ax as (_, consts), old_weight) =>
   439                                         (ax, if exists is_dirty consts then NONE
   440                                              else SOME old_weight)))
   441               val thres =
   442                 1.0 - (1.0 - thres)
   443                       * Math.pow (decay, Real.fromInt (length accepts))
   444               val remaining_max = remaining_max - length accepts
   445             in
   446               trace_msg ctxt (fn () => "New or updated constants: " ^
   447                   commas (rel_const_tab' |> Symtab.dest
   448                           |> subtract (op =) (rel_const_tab |> Symtab.dest)
   449                           |> map string_for_hyper_pconst));
   450               map (fst o fst) accepts @
   451               (if remaining_max = 0 then
   452                  []
   453                else
   454                  iter (j + 1) remaining_max thres rel_const_tab'
   455                       hopeless_rejects hopeful_rejects)
   456             end
   457           | relevant candidates rejects
   458                      (((ax as (((_, stature), _), fact_consts)), cached_weight)
   459                       :: hopeful) =
   460             let
   461               val weight =
   462                 case cached_weight of
   463                   SOME w => w
   464                 | NONE => fact_weight fudge stature const_tab rel_const_tab
   465                                       chained_const_tab fact_consts
   466             in
   467               if weight >= thres then
   468                 relevant ((ax, weight) :: candidates) rejects hopeful
   469               else
   470                 relevant candidates ((ax, weight) :: rejects) hopeful
   471             end
   472         in
   473           trace_msg ctxt (fn () =>
   474               "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
   475               Real.toString thres ^ ", constants: " ^
   476               commas (rel_const_tab |> Symtab.dest
   477                       |> filter (curry (op <>) [] o snd)
   478                       |> map string_for_hyper_pconst));
   479           relevant [] [] hopeful
   480         end
   481     fun uses_const s t =
   482       fold_aterms (curry (fn (Const (s', _), false) => s' = s | (_, b) => b)) t
   483                   false
   484     fun uses_const_anywhere accepts s =
   485       exists (uses_const s o prop_of o snd) accepts orelse
   486       exists (uses_const s) (concl_t :: hyp_ts)
   487     fun add_set_const_thms accepts =
   488       exists (uses_const_anywhere accepts) set_consts ? append set_thms
   489     fun insert_into_facts accepts [] = accepts
   490       | insert_into_facts accepts ths =
   491         let
   492           val add = facts |> filter (member Thm.eq_thm_prop ths o snd)
   493           val (bef, after) =
   494             accepts |> filter_out (member Thm.eq_thm_prop ths o snd)
   495                     |> take (max_facts - length add)
   496                     |> chop special_fact_index
   497         in bef @ add @ after end
   498     fun insert_special_facts accepts =
   499        (* FIXME: get rid of "ext" here once it is treated as a helper *)
   500        [] |> could_benefit_from_ext is_built_in_const accepts ? cons @{thm ext}
   501           |> add_set_const_thms accepts
   502           |> insert_into_facts accepts
   503   in
   504     facts |> map_filter (pair_consts_fact thy is_built_in_const fudge)
   505           |> iter 0 max_facts thres0 goal_const_tab []
   506           |> insert_special_facts
   507           |> tap (fn accepts => trace_msg ctxt (fn () =>
   508                       "Total relevant: " ^ string_of_int (length accepts)))
   509   end
   510 
   511 fun iterative_relevant_facts ctxt
   512         ({fact_thresholds = (thres0, thres1), ...} : params) prover
   513         max_facts fudge hyp_ts concl_t facts =
   514   let
   515     val thy = Proof_Context.theory_of ctxt
   516     val is_built_in_const =
   517       Sledgehammer_Provers.is_built_in_const_for_prover ctxt prover
   518     val fudge =
   519       case fudge of
   520         SOME fudge => fudge
   521       | NONE => Sledgehammer_Provers.relevance_fudge_for_prover ctxt prover
   522     val decay = Math.pow ((1.0 - thres1) / (1.0 - thres0),
   523                           1.0 / Real.fromInt (max_facts + 1))
   524   in
   525     trace_msg ctxt (fn () => "Considering " ^ string_of_int (length facts) ^
   526                              " facts");
   527     (if thres1 < 0.0 then
   528        facts
   529      else if thres0 > 1.0 orelse thres0 > thres1 then
   530        []
   531      else
   532        relevance_filter ctxt thres0 decay max_facts is_built_in_const fudge
   533            facts hyp_ts
   534            (concl_t |> theory_constify fudge (Context.theory_name thy)))
   535   end
   536 
   537 end;