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