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