src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
author blanchet
Wed, 18 Jul 2012 08:44:03 +0200
changeset 49303 255c6e1fd505
parent 49266 6cdcfbddc077
child 49304 6b65f1ad0e4b
permissions -rw-r--r--
rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet@49263
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
blanchet@49263
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@49263
     3
blanchet@49263
     4
Sledgehammer's machine-learning-based relevance filter (MaSh).
blanchet@49263
     5
*)
blanchet@49263
     6
blanchet@49263
     7
signature SLEDGEHAMMER_FILTER_MASH =
blanchet@49263
     8
sig
blanchet@49266
     9
  type status = ATP_Problem_Generate.status
blanchet@49266
    10
  type stature = ATP_Problem_Generate.stature
blanchet@49266
    11
  type params = Sledgehammer_Provers.params
blanchet@49303
    12
  type relevance_override = Sledgehammer_Fact.relevance_override
blanchet@49303
    13
  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
blanchet@49266
    14
  type prover_result = Sledgehammer_Provers.prover_result
blanchet@49266
    15
blanchet@49266
    16
  val fact_name_of : string -> string
blanchet@49266
    17
  val all_non_tautological_facts_of :
blanchet@49266
    18
    theory -> (((unit -> string) * stature) * thm) list
blanchet@49266
    19
  val theory_ord : theory * theory -> order
blanchet@49266
    20
  val thm_ord : thm * thm -> order
blanchet@49266
    21
  val thms_by_thy : ('a * thm) list -> (string * thm list) list
blanchet@49266
    22
  val has_thy : theory -> thm -> bool
blanchet@49266
    23
  val parent_thms : (string * thm list) list -> theory -> string list
blanchet@49266
    24
  val features_of : theory -> status * thm -> string list
blanchet@49266
    25
  val isabelle_dependencies_of : string list -> thm -> string list
blanchet@49266
    26
  val goal_of_thm : theory -> thm -> thm
blanchet@49266
    27
  val iter_facts :
blanchet@49266
    28
    Proof.context -> params -> int -> thm
blanchet@49266
    29
    -> (((unit -> string) * stature) * thm) list
blanchet@49266
    30
    -> ((string * stature) * thm) list
blanchet@49266
    31
  val run_prover :
blanchet@49266
    32
    Proof.context -> params -> ((string * stature) * thm) list -> thm
blanchet@49266
    33
    -> prover_result
blanchet@49266
    34
  val generate_accessibility : theory -> bool -> string -> unit
blanchet@49266
    35
  val generate_features : theory -> bool -> string -> unit
blanchet@49266
    36
  val generate_isa_dependencies : theory -> bool -> string -> unit
blanchet@49266
    37
  val generate_atp_dependencies :
blanchet@49266
    38
    Proof.context -> params -> theory -> bool -> string -> unit
blanchet@49266
    39
blanchet@49264
    40
  val reset : unit -> unit
blanchet@49303
    41
  val can_suggest_facts : unit -> bool
blanchet@49303
    42
(* ###  val suggest_facts : ... *)
blanchet@49264
    43
  val can_learn_thy : theory -> bool
blanchet@49264
    44
  val learn_thy : theory -> real -> unit
blanchet@49264
    45
  val learn_proof : theory -> term -> thm list -> unit
blanchet@49303
    46
blanchet@49303
    47
  val relevant_facts :
blanchet@49303
    48
    Proof.context -> params -> string -> int -> relevance_override -> thm list
blanchet@49303
    49
    -> term list -> term -> (((unit -> string) * stature) * thm) list
blanchet@49303
    50
    -> ((string * stature) * thm) list
blanchet@49263
    51
end;
blanchet@49263
    52
blanchet@49263
    53
structure Sledgehammer_Filter_MaSh : SLEDGEHAMMER_FILTER_MASH =
blanchet@49263
    54
struct
blanchet@49264
    55
blanchet@49266
    56
open ATP_Util
blanchet@49266
    57
open ATP_Problem_Generate
blanchet@49266
    58
open Sledgehammer_Util
blanchet@49266
    59
open Sledgehammer_Fact
blanchet@49266
    60
open Sledgehammer_Filter_Iter
blanchet@49266
    61
open Sledgehammer_Provers
blanchet@49266
    62
blanchet@49266
    63
blanchet@49266
    64
(*** Low-level communication with MaSh ***)
blanchet@49266
    65
blanchet@49264
    66
fun mash_reset () =
blanchet@49264
    67
  tracing "MaSh RESET"
blanchet@49264
    68
blanchet@49264
    69
fun mash_add fact (access, feats, deps) =
blanchet@49264
    70
  tracing ("MaSh ADD " ^ fact ^ ": " ^ space_implode " " access ^ "; " ^
blanchet@49264
    71
           space_implode " " feats ^ "; " ^ space_implode " " deps)
blanchet@49264
    72
blanchet@49264
    73
fun mash_del fact =
blanchet@49264
    74
  tracing ("MaSh DEL " ^ fact)
blanchet@49264
    75
blanchet@49264
    76
fun mash_suggest fact (access, feats) =
blanchet@49264
    77
  tracing ("MaSh SUGGEST " ^ fact ^ ": " ^ space_implode " " access ^ "; " ^
blanchet@49264
    78
           space_implode " " feats)
blanchet@49264
    79
blanchet@49266
    80
blanchet@49266
    81
(*** Isabelle helpers ***)
blanchet@49266
    82
blanchet@49266
    83
val fact_name_of = prefix fact_prefix o ascii_of
blanchet@49266
    84
blanchet@49266
    85
fun escape_meta_char c =
blanchet@49266
    86
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
blanchet@49266
    87
     c = #")" orelse c = #"," then
blanchet@49266
    88
    String.str c
blanchet@49266
    89
  else
blanchet@49266
    90
    (* fixed width, in case more digits follow *)
blanchet@49266
    91
    "\\" ^ stringN_of_int 3 (Char.ord c)
blanchet@49266
    92
blanchet@49266
    93
val escape_meta = String.translate escape_meta_char
blanchet@49266
    94
blanchet@49266
    95
val thy_prefix = "y_"
blanchet@49266
    96
blanchet@49266
    97
val fact_name_of = escape_meta
blanchet@49266
    98
val thy_name_of = prefix thy_prefix o escape_meta
blanchet@49266
    99
val const_name_of = prefix const_prefix o escape_meta
blanchet@49266
   100
val type_name_of = prefix type_const_prefix o escape_meta
blanchet@49266
   101
val class_name_of = prefix class_prefix o escape_meta
blanchet@49266
   102
blanchet@49266
   103
val thy_name_of_thm = theory_of_thm #> Context.theory_name
blanchet@49266
   104
blanchet@49266
   105
local
blanchet@49266
   106
blanchet@49266
   107
fun has_bool @{typ bool} = true
blanchet@49266
   108
  | has_bool (Type (_, Ts)) = exists has_bool Ts
blanchet@49266
   109
  | has_bool _ = false
blanchet@49266
   110
blanchet@49266
   111
fun has_fun (Type (@{type_name fun}, _)) = true
blanchet@49266
   112
  | has_fun (Type (_, Ts)) = exists has_fun Ts
blanchet@49266
   113
  | has_fun _ = false
blanchet@49266
   114
blanchet@49266
   115
val is_conn = member (op =)
blanchet@49266
   116
  [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
blanchet@49266
   117
   @{const_name HOL.implies}, @{const_name Not},
blanchet@49266
   118
   @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex},
blanchet@49266
   119
   @{const_name HOL.eq}]
blanchet@49266
   120
blanchet@49266
   121
val has_bool_arg_const =
blanchet@49266
   122
  exists_Const (fn (c, T) =>
blanchet@49266
   123
                   not (is_conn c) andalso exists has_bool (binder_types T))
blanchet@49266
   124
blanchet@49266
   125
fun higher_inst_const thy (c, T) =
blanchet@49266
   126
  case binder_types T of
blanchet@49266
   127
    [] => false
blanchet@49266
   128
  | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts
blanchet@49266
   129
blanchet@49266
   130
val binders = [@{const_name All}, @{const_name Ex}]
blanchet@49266
   131
blanchet@49266
   132
in
blanchet@49266
   133
blanchet@49266
   134
fun is_fo_term thy t =
blanchet@49266
   135
  let
blanchet@49266
   136
    val t =
blanchet@49266
   137
      t |> Envir.beta_eta_contract
blanchet@49266
   138
        |> transform_elim_prop
blanchet@49266
   139
        |> Object_Logic.atomize_term thy
blanchet@49266
   140
  in
blanchet@49266
   141
    Term.is_first_order binders t andalso
blanchet@49266
   142
    not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T
blanchet@49266
   143
                          | _ => false) t orelse
blanchet@49266
   144
         has_bool_arg_const t orelse exists_Const (higher_inst_const thy) t)
blanchet@49266
   145
  end
blanchet@49266
   146
blanchet@49266
   147
end
blanchet@49266
   148
blanchet@49266
   149
fun interesting_terms_types_and_classes term_max_depth type_max_depth t =
blanchet@49266
   150
  let
blanchet@49266
   151
    val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49266
   152
    val bad_consts = atp_widely_irrelevant_consts
blanchet@49266
   153
    fun do_add_type (Type (s, Ts)) =
blanchet@49266
   154
        (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
blanchet@49266
   155
        #> fold do_add_type Ts
blanchet@49266
   156
      | do_add_type (TFree (_, S)) = union (op =) (map class_name_of S)
blanchet@49266
   157
      | do_add_type (TVar (_, S)) = union (op =) (map class_name_of S)
blanchet@49266
   158
    fun add_type T = type_max_depth >= 0 ? do_add_type T
blanchet@49266
   159
    fun mk_app s args =
blanchet@49266
   160
      if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
blanchet@49266
   161
      else s
blanchet@49266
   162
    fun patternify ~1 _ = ""
blanchet@49266
   163
      | patternify depth t =
blanchet@49266
   164
        case strip_comb t of
blanchet@49266
   165
          (Const (s, _), args) =>
blanchet@49266
   166
          mk_app (const_name_of s) (map (patternify (depth - 1)) args)
blanchet@49266
   167
        | _ => ""
blanchet@49266
   168
    fun add_term_patterns ~1 _ = I
blanchet@49266
   169
      | add_term_patterns depth t =
blanchet@49266
   170
        insert (op =) (patternify depth t)
blanchet@49266
   171
        #> add_term_patterns (depth - 1) t
blanchet@49266
   172
    val add_term = add_term_patterns term_max_depth
blanchet@49266
   173
    fun add_patterns t =
blanchet@49266
   174
      let val (head, args) = strip_comb t in
blanchet@49266
   175
        (case head of
blanchet@49266
   176
           Const (s, T) =>
blanchet@49266
   177
           not (member (op =) bad_consts s) ? (add_term t #> add_type T)
blanchet@49266
   178
         | Free (_, T) => add_type T
blanchet@49266
   179
         | Var (_, T) => add_type T
blanchet@49266
   180
         | Abs (_, T, body) => add_type T #> add_patterns body
blanchet@49266
   181
         | _ => I)
blanchet@49266
   182
        #> fold add_patterns args
blanchet@49266
   183
      end
blanchet@49266
   184
  in [] |> add_patterns t |> sort string_ord end
blanchet@49266
   185
blanchet@49266
   186
fun is_likely_tautology th =
blanchet@49266
   187
  null (interesting_terms_types_and_classes 0 ~1 (prop_of th)) andalso
blanchet@49266
   188
  not (Thm.eq_thm_prop (@{thm ext}, th))
blanchet@49266
   189
blanchet@49266
   190
fun is_too_meta thy th =
blanchet@49266
   191
  fastype_of (Object_Logic.atomize_term thy (prop_of th)) <> @{typ bool}
blanchet@49266
   192
blanchet@49266
   193
fun all_non_tautological_facts_of thy =
blanchet@49266
   194
  all_facts_of thy
blanchet@49266
   195
  |> filter_out ((is_likely_tautology orf is_too_meta thy) o snd)
blanchet@49266
   196
blanchet@49266
   197
fun theory_ord p =
blanchet@49266
   198
  if Theory.eq_thy p then EQUAL
blanchet@49266
   199
  else if Theory.subthy p then LESS
blanchet@49266
   200
  else if Theory.subthy (swap p) then GREATER
blanchet@49266
   201
  else EQUAL
blanchet@49266
   202
blanchet@49266
   203
val thm_ord = theory_ord o pairself theory_of_thm
blanchet@49266
   204
blanchet@49266
   205
fun thms_by_thy ths =
blanchet@49266
   206
  ths |> map (snd #> `thy_name_of_thm)
blanchet@49266
   207
      |> AList.group (op =)
blanchet@49266
   208
      |> sort (int_ord o pairself (length o Theory.ancestors_of o theory_of_thm
blanchet@49266
   209
                                   o hd o snd))
blanchet@49266
   210
      |> map (apsnd (sort thm_ord))
blanchet@49266
   211
blanchet@49266
   212
fun has_thy thy th = (Context.theory_name thy = thy_name_of_thm th)
blanchet@49266
   213
blanchet@49266
   214
fun parent_thms thy_ths thy =
blanchet@49266
   215
  Theory.parents_of thy
blanchet@49266
   216
  |> map Context.theory_name
blanchet@49266
   217
  |> map_filter (AList.lookup (op =) thy_ths)
blanchet@49266
   218
  |> map List.last
blanchet@49266
   219
  |> map (fact_name_of o Thm.get_name_hint)
blanchet@49266
   220
blanchet@49266
   221
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
blanchet@49266
   222
blanchet@49266
   223
val max_depth = 1
blanchet@49266
   224
blanchet@49266
   225
(* TODO: Generate type classes for types? *)
blanchet@49266
   226
fun features_of thy (status, th) =
blanchet@49266
   227
  let val t = Thm.prop_of th in
blanchet@49266
   228
    thy_name_of (thy_name_of_thm th) ::
blanchet@49266
   229
    interesting_terms_types_and_classes max_depth max_depth t
blanchet@49266
   230
    |> not (has_no_lambdas t) ? cons "lambdas"
blanchet@49266
   231
    |> exists_Const is_exists t ? cons "skolems"
blanchet@49266
   232
    |> not (is_fo_term thy t) ? cons "ho"
blanchet@49266
   233
    |> (case status of
blanchet@49266
   234
          General => I
blanchet@49266
   235
        | Induction => cons "induction"
blanchet@49266
   236
        | Intro => cons "intro"
blanchet@49266
   237
        | Inductive => cons "inductive"
blanchet@49266
   238
        | Elim => cons "elim"
blanchet@49266
   239
        | Simp => cons "simp"
blanchet@49266
   240
        | Def => cons "def")
blanchet@49266
   241
  end
blanchet@49266
   242
blanchet@49266
   243
fun isabelle_dependencies_of all_facts =
blanchet@49266
   244
  thms_in_proof (SOME all_facts)
blanchet@49266
   245
  #> map fact_name_of #> sort string_ord
blanchet@49266
   246
blanchet@49266
   247
val freezeT = Type.legacy_freeze_type
blanchet@49266
   248
blanchet@49266
   249
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49266
   250
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49266
   251
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49266
   252
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49266
   253
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49266
   254
  | freeze t = t
blanchet@49266
   255
blanchet@49266
   256
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
blanchet@49266
   257
blanchet@49303
   258
fun iter_facts ctxt (params as {provers, ...}) max_relevant goal =
blanchet@49303
   259
  let val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1 in
blanchet@49303
   260
    iterative_relevant_facts ctxt params (hd provers) max_relevant NONE
blanchet@49303
   261
                             no_relevance_override [] hyp_ts concl_t
blanchet@49266
   262
  end
blanchet@49266
   263
blanchet@49266
   264
fun run_prover ctxt (params as {provers, ...}) facts goal =
blanchet@49266
   265
  let
blanchet@49266
   266
    val problem =
blanchet@49266
   267
      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
blanchet@49266
   268
       facts = facts |> map Sledgehammer_Provers.Untranslated_Fact}
blanchet@49266
   269
    val prover =
blanchet@49266
   270
      Sledgehammer_Minimize.get_minimizing_prover ctxt
blanchet@49266
   271
          Sledgehammer_Provers.Normal (hd provers)
blanchet@49266
   272
  in prover params (K (K (K ""))) problem end
blanchet@49266
   273
blanchet@49266
   274
fun generate_accessibility thy include_thy file_name =
blanchet@49266
   275
  let
blanchet@49266
   276
    val path = file_name |> Path.explode
blanchet@49266
   277
    val _ = File.write path ""
blanchet@49266
   278
    fun do_thm th prevs =
blanchet@49266
   279
      let
blanchet@49266
   280
        val s = th ^ ": " ^ space_implode " " prevs ^ "\n"
blanchet@49266
   281
        val _ = File.append path s
blanchet@49266
   282
      in [th] end
blanchet@49266
   283
    val thy_ths =
blanchet@49266
   284
      all_non_tautological_facts_of thy
blanchet@49266
   285
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49266
   286
      |> thms_by_thy
blanchet@49266
   287
    fun do_thy ths =
blanchet@49266
   288
      let
blanchet@49266
   289
        val thy = theory_of_thm (hd ths)
blanchet@49266
   290
        val parents = parent_thms thy_ths thy
blanchet@49266
   291
        val ths = ths |> map (fact_name_of o Thm.get_name_hint)
blanchet@49266
   292
      in fold do_thm ths parents; () end
blanchet@49266
   293
  in List.app (do_thy o snd) thy_ths end
blanchet@49266
   294
blanchet@49266
   295
fun generate_features thy include_thy file_name =
blanchet@49266
   296
  let
blanchet@49266
   297
    val path = file_name |> Path.explode
blanchet@49266
   298
    val _ = File.write path ""
blanchet@49266
   299
    val facts =
blanchet@49266
   300
      all_non_tautological_facts_of thy
blanchet@49266
   301
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49266
   302
    fun do_fact ((_, (_, status)), th) =
blanchet@49266
   303
      let
blanchet@49266
   304
        val name = Thm.get_name_hint th
blanchet@49266
   305
        val feats = features_of thy (status, th)
blanchet@49266
   306
        val s = fact_name_of name ^ ": " ^ space_implode " " feats ^ "\n"
blanchet@49266
   307
      in File.append path s end
blanchet@49266
   308
  in List.app do_fact facts end
blanchet@49266
   309
blanchet@49266
   310
fun generate_isa_dependencies thy include_thy file_name =
blanchet@49266
   311
  let
blanchet@49266
   312
    val path = file_name |> Path.explode
blanchet@49266
   313
    val _ = File.write path ""
blanchet@49266
   314
    val ths =
blanchet@49266
   315
      all_non_tautological_facts_of thy
blanchet@49266
   316
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49266
   317
      |> map snd
blanchet@49266
   318
    val all_names = ths |> map Thm.get_name_hint
blanchet@49266
   319
    fun do_thm th =
blanchet@49266
   320
      let
blanchet@49266
   321
        val name = Thm.get_name_hint th
blanchet@49266
   322
        val deps = isabelle_dependencies_of all_names th
blanchet@49266
   323
        val s = fact_name_of name ^ ": " ^ space_implode " " deps ^ "\n"
blanchet@49266
   324
      in File.append path s end
blanchet@49266
   325
  in List.app do_thm ths end
blanchet@49266
   326
blanchet@49266
   327
fun generate_atp_dependencies ctxt (params as {max_relevant, ...}) thy
blanchet@49266
   328
                              include_thy file_name =
blanchet@49266
   329
  let
blanchet@49266
   330
    val path = file_name |> Path.explode
blanchet@49266
   331
    val _ = File.write path ""
blanchet@49266
   332
    val facts =
blanchet@49266
   333
      all_non_tautological_facts_of thy
blanchet@49266
   334
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49266
   335
    val ths = facts |> map snd
blanchet@49266
   336
    val all_names = ths |> map Thm.get_name_hint
blanchet@49266
   337
    fun is_dep dep (_, th) = fact_name_of (Thm.get_name_hint th) = dep
blanchet@49266
   338
    fun add_isa_dep facts dep accum =
blanchet@49266
   339
      if exists (is_dep dep) accum then
blanchet@49266
   340
        accum
blanchet@49266
   341
      else case find_first (is_dep dep) facts of
blanchet@49266
   342
        SOME ((name, status), th) => accum @ [((name (), status), th)]
blanchet@49266
   343
      | NONE => accum (* shouldn't happen *)
blanchet@49266
   344
    fun fix_name ((_, stature), th) =
blanchet@49266
   345
      ((th |> Thm.get_name_hint |> fact_name_of, stature), th)
blanchet@49266
   346
    fun do_thm th =
blanchet@49266
   347
      let
blanchet@49266
   348
        val name = Thm.get_name_hint th
blanchet@49266
   349
        val goal = goal_of_thm thy th
blanchet@49266
   350
        val deps =
blanchet@49266
   351
          case isabelle_dependencies_of all_names th of
blanchet@49266
   352
            [] => []
blanchet@49266
   353
          | isa_dep as [_] => isa_dep (* can hardly beat that *)
blanchet@49266
   354
          | isa_deps =>
blanchet@49266
   355
            let
blanchet@49266
   356
              val facts =
blanchet@49266
   357
                facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
blanchet@49266
   358
              val facts =
blanchet@49266
   359
                facts |> iter_facts ctxt params (the max_relevant) goal
blanchet@49266
   360
                      |> fold (add_isa_dep facts) isa_deps
blanchet@49266
   361
                      |> map fix_name
blanchet@49266
   362
            in
blanchet@49266
   363
              case run_prover ctxt params facts goal of
blanchet@49266
   364
                {outcome = NONE, used_facts, ...} =>
blanchet@49266
   365
                used_facts |> map fst |> sort string_ord
blanchet@49266
   366
              | _ => isa_deps
blanchet@49266
   367
            end
blanchet@49266
   368
        val s = fact_name_of name ^ ": " ^ space_implode " " deps ^ "\n"
blanchet@49266
   369
      in File.append path s end
blanchet@49266
   370
  in List.app do_thm ths end
blanchet@49266
   371
blanchet@49266
   372
blanchet@49266
   373
(*** High-level communication with MaSh ***)
blanchet@49266
   374
blanchet@49264
   375
fun reset () =
blanchet@49264
   376
  ()
blanchet@49264
   377
blanchet@49303
   378
fun can_suggest_facts () =
blanchet@49264
   379
  true
blanchet@49264
   380
blanchet@49264
   381
fun can_learn_thy thy =
blanchet@49264
   382
  true
blanchet@49264
   383
blanchet@49264
   384
fun learn_thy thy timeout =
blanchet@49264
   385
  ()
blanchet@49264
   386
blanchet@49303
   387
fun learn_proof thy t thms =
blanchet@49264
   388
  ()
blanchet@49264
   389
blanchet@49303
   390
fun relevant_facts ctxt params prover max_relevant
blanchet@49303
   391
                   (override as {add, only, ...}) chained_ths hyp_ts concl_t
blanchet@49303
   392
                   facts =
blanchet@49303
   393
  if only then
blanchet@49303
   394
    facts |> map (apfst (apfst (fn f => f ()))) (* ###*)
blanchet@49303
   395
  else if max_relevant <= 0 then
blanchet@49303
   396
    []
blanchet@49303
   397
  else
blanchet@49303
   398
    let
blanchet@49303
   399
      val add_ths = Attrib.eval_thms ctxt add
blanchet@49303
   400
      fun prepend_facts ths facts =
blanchet@49303
   401
        ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
blanchet@49303
   402
         (facts |> filter_out (member Thm.eq_thm_prop ths o snd)))
blanchet@49303
   403
        |> take max_relevant
blanchet@49303
   404
    in
blanchet@49303
   405
      iterative_relevant_facts ctxt params prover max_relevant NONE override
blanchet@49303
   406
                               chained_ths hyp_ts concl_t facts
blanchet@49303
   407
      |> not (null add_ths) ? prepend_facts add_ths
blanchet@49303
   408
    end
blanchet@49303
   409
blanchet@49263
   410
end;