src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML
author blanchet
Fri, 20 Jul 2012 22:19:46 +0200
changeset 49409 82fc8c956cdc
parent 49407 ca998fa08cd9
child 49410 85a7fb65507a
permissions -rw-r--r--
fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
blanchet@49395
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_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@49396
     7
signature SLEDGEHAMMER_MASH =
blanchet@49263
     8
sig
blanchet@49266
     9
  type stature = ATP_Problem_Generate.stature
blanchet@49311
    10
  type fact = Sledgehammer_Fact.fact
blanchet@49311
    11
  type fact_override = Sledgehammer_Fact.fact_override
blanchet@49266
    12
  type params = Sledgehammer_Provers.params
blanchet@49303
    13
  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
blanchet@49266
    14
  type prover_result = Sledgehammer_Provers.prover_result
blanchet@49266
    15
blanchet@49323
    16
  val trace : bool Config.T
blanchet@49334
    17
  val MaShN : string
blanchet@49394
    18
  val mepoN : string
blanchet@49394
    19
  val mashN : string
blanchet@49329
    20
  val meshN : string
blanchet@49407
    21
  val unlearnN : string
blanchet@49407
    22
  val learn_isarN : string
blanchet@49407
    23
  val learn_atpN : string
blanchet@49407
    24
  val relearn_isarN : string
blanchet@49407
    25
  val relearn_atpN : string
blanchet@49329
    26
  val fact_filters : string list
blanchet@49318
    27
  val escape_meta : string -> string
blanchet@49318
    28
  val escape_metas : string list -> string
blanchet@49323
    29
  val unescape_meta : string -> string
blanchet@49323
    30
  val unescape_metas : string -> string list
blanchet@49326
    31
  val extract_query : string -> string * string list
blanchet@49393
    32
  val nickname_of : thm -> string
blanchet@49336
    33
  val suggested_facts : string list -> ('a * thm) list -> ('a * thm) list
blanchet@49336
    34
  val mesh_facts :
blanchet@49336
    35
    int -> (('a * thm) list * ('a * thm) list) list -> ('a * thm) list
blanchet@49339
    36
  val is_likely_tautology_or_too_meta : thm -> bool
blanchet@49266
    37
  val theory_ord : theory * theory -> order
blanchet@49266
    38
  val thm_ord : thm * thm -> order
blanchet@49266
    39
  val goal_of_thm : theory -> thm -> thm
blanchet@49336
    40
  val run_prover_for_mash :
blanchet@49333
    41
    Proof.context -> params -> string -> fact list -> thm -> prover_result
blanchet@49407
    42
  val features_of :
blanchet@49407
    43
    Proof.context -> string -> theory -> stature -> term list -> string list
blanchet@49407
    44
  val isar_dependencies_of : unit Symtab.table -> thm -> string list
blanchet@49407
    45
  val atp_dependencies_of :
blanchet@49407
    46
    Proof.context -> params -> string -> bool -> fact list -> unit Symtab.table
blanchet@49407
    47
    -> thm -> string list
blanchet@49347
    48
  val mash_CLEAR : Proof.context -> unit
blanchet@49323
    49
  val mash_ADD :
blanchet@49331
    50
    Proof.context -> bool
blanchet@49331
    51
    -> (string * string list * string list * string list) list -> unit
blanchet@49331
    52
  val mash_QUERY :
blanchet@49333
    53
    Proof.context -> bool -> int -> string list * string list -> string list
blanchet@49347
    54
  val mash_unlearn : Proof.context -> unit
blanchet@49333
    55
  val mash_could_suggest_facts : unit -> bool
blanchet@49336
    56
  val mash_can_suggest_facts : Proof.context -> bool
blanchet@49313
    57
  val mash_suggest_facts :
blanchet@49336
    58
    Proof.context -> params -> string -> int -> term list -> term
blanchet@49336
    59
    -> ('a * thm) list -> ('a * thm) list * ('a * thm) list
blanchet@49398
    60
  val mash_learn_proof :
blanchet@49399
    61
    Proof.context -> params -> string -> term -> ('a * thm) list -> thm list
blanchet@49399
    62
    -> unit
blanchet@49405
    63
  val mash_learn_facts :
blanchet@49407
    64
    Proof.context -> params -> string -> bool -> bool -> Time.time -> fact list
blanchet@49407
    65
    -> string
blanchet@49407
    66
  val mash_learn : Proof.context -> params -> bool -> unit
blanchet@49303
    67
  val relevant_facts :
blanchet@49307
    68
    Proof.context -> params -> string -> int -> fact_override -> term list
blanchet@49311
    69
    -> term -> fact list -> fact list
blanchet@49334
    70
  val kill_learners : unit -> unit
blanchet@49334
    71
  val running_learners : unit -> unit
blanchet@49263
    72
end;
blanchet@49263
    73
blanchet@49396
    74
structure Sledgehammer_MaSh : SLEDGEHAMMER_MASH =
blanchet@49263
    75
struct
blanchet@49264
    76
blanchet@49266
    77
open ATP_Util
blanchet@49266
    78
open ATP_Problem_Generate
blanchet@49266
    79
open Sledgehammer_Util
blanchet@49266
    80
open Sledgehammer_Fact
blanchet@49266
    81
open Sledgehammer_Provers
blanchet@49333
    82
open Sledgehammer_Minimize
blanchet@49396
    83
open Sledgehammer_MePo
blanchet@49266
    84
blanchet@49323
    85
val trace =
blanchet@49395
    86
  Attrib.setup_config_bool @{binding sledgehammer_mash_trace} (K false)
blanchet@49323
    87
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
blanchet@49323
    88
blanchet@49334
    89
val MaShN = "MaSh"
blanchet@49334
    90
blanchet@49394
    91
val mepoN = "mepo"
blanchet@49394
    92
val mashN = "mash"
blanchet@49329
    93
val meshN = "mesh"
blanchet@49329
    94
blanchet@49394
    95
val fact_filters = [meshN, mepoN, mashN]
blanchet@49329
    96
blanchet@49407
    97
val unlearnN = "unlearn"
blanchet@49407
    98
val learn_isarN = "learn_isar"
blanchet@49407
    99
val learn_atpN = "learn_atp"
blanchet@49407
   100
val relearn_isarN = "relearn_isar"
blanchet@49407
   101
val relearn_atpN = "relearn_atp"
blanchet@49407
   102
blanchet@49329
   103
fun mash_home () = getenv "MASH_HOME"
blanchet@49409
   104
fun mash_model_dir () =
blanchet@49324
   105
  getenv "ISABELLE_HOME_USER" ^ "/mash"
blanchet@49331
   106
  |> tap (Isabelle_System.mkdir o Path.explode)
blanchet@49409
   107
val mash_state_dir = mash_model_dir
blanchet@49329
   108
fun mash_state_path () = mash_state_dir () ^ "/state" |> Path.explode
blanchet@49266
   109
blanchet@49345
   110
blanchet@49266
   111
(*** Isabelle helpers ***)
blanchet@49266
   112
blanchet@49323
   113
fun meta_char c =
blanchet@49266
   114
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
blanchet@49266
   115
     c = #")" orelse c = #"," then
blanchet@49266
   116
    String.str c
blanchet@49266
   117
  else
blanchet@49266
   118
    (* fixed width, in case more digits follow *)
blanchet@49409
   119
    "$" ^ stringN_of_int 3 (Char.ord c)
blanchet@49266
   120
blanchet@49323
   121
fun unmeta_chars accum [] = String.implode (rev accum)
blanchet@49409
   122
  | unmeta_chars accum (#"$" :: d1 :: d2 :: d3 :: cs) =
blanchet@49323
   123
    (case Int.fromString (String.implode [d1, d2, d3]) of
blanchet@49323
   124
       SOME n => unmeta_chars (Char.chr n :: accum) cs
blanchet@49323
   125
     | NONE => "" (* error *))
blanchet@49409
   126
  | unmeta_chars _ (#"$" :: _) = "" (* error *)
blanchet@49323
   127
  | unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs
blanchet@49323
   128
blanchet@49323
   129
val escape_meta = String.translate meta_char
blanchet@49318
   130
val escape_metas = map escape_meta #> space_implode " "
blanchet@49330
   131
val unescape_meta = String.explode #> unmeta_chars []
blanchet@49330
   132
val unescape_metas =
blanchet@49330
   133
  space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta
blanchet@49266
   134
blanchet@49326
   135
fun extract_query line =
blanchet@49326
   136
  case space_explode ":" line of
blanchet@49330
   137
    [goal_name, suggs] => (unescape_meta goal_name, unescape_metas suggs)
blanchet@49327
   138
  | _ => ("", [])
blanchet@49326
   139
blanchet@49393
   140
fun parent_of_local_thm th =
blanchet@49393
   141
  let
blanchet@49393
   142
    val thy = th |> Thm.theory_of_thm
blanchet@49393
   143
    val facts = thy |> Global_Theory.facts_of
blanchet@49393
   144
    val space = facts |> Facts.space_of
blanchet@49393
   145
    fun id_of s = #id (Name_Space.the_entry space s)
blanchet@49393
   146
    fun max_id (s', _) (s, id) =
blanchet@49393
   147
      let val id' = id_of s' in if id > id' then (s, id) else (s', id') end
blanchet@49393
   148
  in ("", ~1) |> Facts.fold_static max_id facts |> fst end
blanchet@49393
   149
blanchet@49393
   150
val local_prefix = "local" ^ Long_Name.separator
blanchet@49393
   151
blanchet@49393
   152
fun nickname_of th =
blanchet@49409
   153
  if Thm.has_name_hint th then
blanchet@49409
   154
    let val hint = Thm.get_name_hint th in
blanchet@49409
   155
      (* FIXME: There must be a better way to detect local facts. *)
blanchet@49409
   156
      case try (unprefix local_prefix) hint of
blanchet@49409
   157
        SOME suf =>
blanchet@49409
   158
        parent_of_local_thm th ^ Long_Name.separator ^ Long_Name.separator ^ suf
blanchet@49409
   159
      | NONE => hint
blanchet@49409
   160
    end
blanchet@49409
   161
  else
blanchet@49409
   162
    backquote_thm th
blanchet@49393
   163
blanchet@49345
   164
fun suggested_facts suggs facts =
blanchet@49345
   165
  let
blanchet@49393
   166
    fun add_fact (fact as (_, th)) = Symtab.default (nickname_of th, fact)
blanchet@49345
   167
    val tab = Symtab.empty |> fold add_fact facts
blanchet@49345
   168
  in map_filter (Symtab.lookup tab) suggs end
blanchet@49326
   169
blanchet@49344
   170
(* Ad hoc score function roughly based on Blanchette's Ringberg 2011 data. *)
blanchet@49344
   171
fun score x = Math.pow (1.5, 15.5 - 0.05 * Real.fromInt x) + 15.0
blanchet@49343
   172
blanchet@49343
   173
fun sum_sq_avg [] = 0
blanchet@49344
   174
  | sum_sq_avg xs =
blanchet@49344
   175
    Real.ceil (100000.0 * fold (curry (op +) o score) xs 0.0) div length xs
blanchet@49328
   176
blanchet@49335
   177
fun mesh_facts max_facts [(selected, unknown)] =
blanchet@49335
   178
    take max_facts selected @ take (max_facts - length selected) unknown
blanchet@49329
   179
  | mesh_facts max_facts mess =
blanchet@49329
   180
    let
blanchet@49335
   181
      val mess = mess |> map (apfst (`length))
blanchet@49329
   182
      val fact_eq = Thm.eq_thm o pairself snd
blanchet@49335
   183
      fun score_in fact ((sel_len, sels), unks) =
blanchet@49335
   184
        case find_index (curry fact_eq fact) sels of
blanchet@49335
   185
          ~1 => (case find_index (curry fact_eq fact) unks of
blanchet@49344
   186
                   ~1 => SOME sel_len
blanchet@49335
   187
                 | _ => NONE)
blanchet@49344
   188
        | j => SOME j
blanchet@49343
   189
      fun score_of fact = mess |> map_filter (score_in fact) |> sum_sq_avg
blanchet@49335
   190
      val facts = fold (union fact_eq o take max_facts o snd o fst) mess []
blanchet@49329
   191
    in
blanchet@49343
   192
      facts |> map (`score_of) |> sort (int_ord o swap o pairself fst)
blanchet@49343
   193
            |> map snd |> take max_facts
blanchet@49329
   194
    end
blanchet@49327
   195
blanchet@49318
   196
val thy_feature_prefix = "y_"
blanchet@49266
   197
blanchet@49318
   198
val thy_feature_name_of = prefix thy_feature_prefix
blanchet@49318
   199
val const_name_of = prefix const_prefix
blanchet@49318
   200
val type_name_of = prefix type_const_prefix
blanchet@49318
   201
val class_name_of = prefix class_prefix
blanchet@49266
   202
blanchet@49339
   203
fun is_likely_tautology_or_too_meta th =
blanchet@49339
   204
  let
blanchet@49339
   205
    val is_boring_const = member (op =) atp_widely_irrelevant_consts
blanchet@49339
   206
    fun is_boring_bool t =
blanchet@49339
   207
      not (exists_Const (not o is_boring_const o fst) t) orelse
blanchet@49339
   208
      exists_type (exists_subtype (curry (op =) @{typ prop})) t
blanchet@49339
   209
    fun is_boring_prop (@{const Trueprop} $ t) = is_boring_bool t
blanchet@49339
   210
      | is_boring_prop (@{const "==>"} $ t $ u) =
blanchet@49339
   211
        is_boring_prop t andalso is_boring_prop u
blanchet@49339
   212
      | is_boring_prop (Const (@{const_name all}, _) $ (Abs (_, _, t)) $ u) =
blanchet@49339
   213
        is_boring_prop t andalso is_boring_prop u
blanchet@49339
   214
      | is_boring_prop (Const (@{const_name "=="}, _) $ t $ u) =
blanchet@49339
   215
        is_boring_bool t andalso is_boring_bool u
blanchet@49339
   216
      | is_boring_prop _ = true
blanchet@49339
   217
  in
blanchet@49339
   218
    is_boring_prop (prop_of th) andalso not (Thm.eq_thm_prop (@{thm ext}, th))
blanchet@49339
   219
  end
blanchet@49339
   220
blanchet@49339
   221
fun theory_ord p =
blanchet@49339
   222
  if Theory.eq_thy p then
blanchet@49339
   223
    EQUAL
blanchet@49339
   224
  else if Theory.subthy p then
blanchet@49339
   225
    LESS
blanchet@49339
   226
  else if Theory.subthy (swap p) then
blanchet@49339
   227
    GREATER
blanchet@49339
   228
  else case int_ord (pairself (length o Theory.ancestors_of) p) of
blanchet@49339
   229
    EQUAL => string_ord (pairself Context.theory_name p)
blanchet@49339
   230
  | order => order
blanchet@49339
   231
blanchet@49339
   232
val thm_ord = theory_ord o pairself theory_of_thm
blanchet@49339
   233
blanchet@49407
   234
val freezeT = Type.legacy_freeze_type
blanchet@49407
   235
blanchet@49407
   236
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49407
   237
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49407
   238
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49407
   239
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49407
   240
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49407
   241
  | freeze t = t
blanchet@49407
   242
blanchet@49407
   243
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
blanchet@49407
   244
blanchet@49407
   245
fun run_prover_for_mash ctxt params prover facts goal =
blanchet@49407
   246
  let
blanchet@49407
   247
    val problem =
blanchet@49407
   248
      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
blanchet@49407
   249
       facts = facts |> map (apfst (apfst (fn name => name ())))
blanchet@49407
   250
                     |> map Untranslated_Fact}
blanchet@49407
   251
  in
blanchet@49407
   252
    get_minimizing_prover ctxt MaSh (K ()) prover params (K (K (K "")))
blanchet@49407
   253
                          problem
blanchet@49407
   254
  end
blanchet@49407
   255
blanchet@49341
   256
val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49341
   257
blanchet@49333
   258
fun interesting_terms_types_and_classes ctxt prover term_max_depth
blanchet@49333
   259
                                        type_max_depth ts =
blanchet@49266
   260
  let
blanchet@49333
   261
    fun is_bad_const (x as (s, _)) args =
blanchet@49333
   262
      member (op =) atp_logical_consts s orelse
blanchet@49333
   263
      fst (is_built_in_const_for_prover ctxt prover x args)
blanchet@49319
   264
    fun add_classes @{sort type} = I
blanchet@49319
   265
      | add_classes S = union (op =) (map class_name_of S)
blanchet@49266
   266
    fun do_add_type (Type (s, Ts)) =
blanchet@49266
   267
        (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
blanchet@49266
   268
        #> fold do_add_type Ts
blanchet@49319
   269
      | do_add_type (TFree (_, S)) = add_classes S
blanchet@49319
   270
      | do_add_type (TVar (_, S)) = add_classes S
blanchet@49266
   271
    fun add_type T = type_max_depth >= 0 ? do_add_type T
blanchet@49266
   272
    fun mk_app s args =
blanchet@49266
   273
      if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
blanchet@49266
   274
      else s
blanchet@49266
   275
    fun patternify ~1 _ = ""
blanchet@49266
   276
      | patternify depth t =
blanchet@49266
   277
        case strip_comb t of
blanchet@49266
   278
          (Const (s, _), args) =>
blanchet@49266
   279
          mk_app (const_name_of s) (map (patternify (depth - 1)) args)
blanchet@49266
   280
        | _ => ""
blanchet@49266
   281
    fun add_term_patterns ~1 _ = I
blanchet@49266
   282
      | add_term_patterns depth t =
blanchet@49266
   283
        insert (op =) (patternify depth t)
blanchet@49266
   284
        #> add_term_patterns (depth - 1) t
blanchet@49266
   285
    val add_term = add_term_patterns term_max_depth
blanchet@49266
   286
    fun add_patterns t =
blanchet@49266
   287
      let val (head, args) = strip_comb t in
blanchet@49266
   288
        (case head of
blanchet@49333
   289
           Const (x as (_, T)) =>
blanchet@49333
   290
           not (is_bad_const x args) ? (add_term t #> add_type T)
blanchet@49266
   291
         | Free (_, T) => add_type T
blanchet@49266
   292
         | Var (_, T) => add_type T
blanchet@49266
   293
         | Abs (_, T, body) => add_type T #> add_patterns body
blanchet@49266
   294
         | _ => I)
blanchet@49266
   295
        #> fold add_patterns args
blanchet@49266
   296
      end
blanchet@49341
   297
  in [] |> fold add_patterns ts end
blanchet@49266
   298
blanchet@49266
   299
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
blanchet@49266
   300
blanchet@49312
   301
val term_max_depth = 1
blanchet@49312
   302
val type_max_depth = 1
blanchet@49266
   303
blanchet@49266
   304
(* TODO: Generate type classes for types? *)
blanchet@49400
   305
fun features_of ctxt prover thy (scope, status) ts =
blanchet@49318
   306
  thy_feature_name_of (Context.theory_name thy) ::
blanchet@49333
   307
  interesting_terms_types_and_classes ctxt prover term_max_depth type_max_depth
blanchet@49333
   308
                                      ts
blanchet@49347
   309
  |> forall is_lambda_free ts ? cons "no_lams"
blanchet@49347
   310
  |> forall (not o exists_Const is_exists) ts ? cons "no_skos"
blanchet@49400
   311
  |> scope <> Global ? cons "local"
blanchet@49317
   312
  |> (case status of
blanchet@49317
   313
        General => I
blanchet@49317
   314
      | Induction => cons "induction"
blanchet@49317
   315
      | Intro => cons "intro"
blanchet@49317
   316
      | Inductive => cons "inductive"
blanchet@49317
   317
      | Elim => cons "elim"
blanchet@49317
   318
      | Simp => cons "simp"
blanchet@49317
   319
      | Def => cons "def")
blanchet@49266
   320
blanchet@49407
   321
fun isar_dependencies_of all_facts = thms_in_proof (SOME all_facts)
blanchet@49266
   322
blanchet@49407
   323
val atp_dep_default_max_fact = 50
blanchet@49266
   324
blanchet@49407
   325
fun atp_dependencies_of ctxt (params as {verbose, max_facts, ...}) prover auto
blanchet@49407
   326
                        facts all_names th =
blanchet@49407
   327
  case isar_dependencies_of all_names th of
blanchet@49407
   328
    [] => []
blanchet@49407
   329
  | isar_deps =>
blanchet@49407
   330
    let
blanchet@49407
   331
      val thy = Proof_Context.theory_of ctxt
blanchet@49407
   332
      val goal = goal_of_thm thy th
blanchet@49407
   333
      val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
blanchet@49407
   334
      val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
blanchet@49407
   335
      fun fix_name ((_, stature), th) = ((fn () => nickname_of th, stature), th)
blanchet@49407
   336
      fun is_dep dep (_, th) = nickname_of th = dep
blanchet@49407
   337
      fun add_isar_dep facts dep accum =
blanchet@49407
   338
        if exists (is_dep dep) accum then
blanchet@49407
   339
          accum
blanchet@49407
   340
        else case find_first (is_dep dep) facts of
blanchet@49407
   341
          SOME ((name, status), th) => accum @ [((name, status), th)]
blanchet@49407
   342
        | NONE => accum (* shouldn't happen *)
blanchet@49407
   343
      val facts =
blanchet@49407
   344
        facts |> iterative_relevant_facts ctxt params prover
blanchet@49407
   345
                     (max_facts |> the_default atp_dep_default_max_fact) NONE
blanchet@49407
   346
                     hyp_ts concl_t
blanchet@49407
   347
              |> fold (add_isar_dep facts) isar_deps
blanchet@49407
   348
              |> map fix_name
blanchet@49407
   349
    in
blanchet@49407
   350
      if verbose andalso not auto then
blanchet@49407
   351
        let val num_facts = length facts in
blanchet@49407
   352
          "MaSh: " ^ quote prover ^ " on " ^ quote (nickname_of th) ^
blanchet@49407
   353
          " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
blanchet@49407
   354
          "."
blanchet@49407
   355
          |> Output.urgent_message
blanchet@49407
   356
        end
blanchet@49407
   357
      else
blanchet@49407
   358
        ();
blanchet@49407
   359
      case run_prover_for_mash ctxt params prover facts goal of
blanchet@49407
   360
        {outcome = NONE, used_facts, ...} =>
blanchet@49407
   361
        (if verbose andalso not auto then
blanchet@49407
   362
           let val num_facts = length used_facts in
blanchet@49407
   363
             "Found proof with " ^ string_of_int num_facts ^ " fact" ^
blanchet@49407
   364
             plural_s num_facts ^ "."
blanchet@49407
   365
             |> Output.urgent_message
blanchet@49407
   366
           end
blanchet@49407
   367
         else
blanchet@49407
   368
           ();
blanchet@49407
   369
         used_facts |> map fst)
blanchet@49407
   370
      | _ => isar_deps
blanchet@49407
   371
    end
blanchet@49266
   372
blanchet@49266
   373
blanchet@49317
   374
(*** Low-level communication with MaSh ***)
blanchet@49317
   375
blanchet@49409
   376
(* more friendly than "try o File.rm" for those who keep the files open in their
blanchet@49409
   377
   text editor *)
blanchet@49409
   378
fun wipe_out file = File.write file ""
blanchet@49409
   379
blanchet@49338
   380
fun write_file (xs, f) file =
blanchet@49333
   381
  let val path = Path.explode file in
blanchet@49409
   382
    wipe_out path;
blanchet@49338
   383
    xs |> chunk_list 500
blanchet@49338
   384
       |> List.app (File.append path o space_implode "" o map f)
blanchet@49333
   385
  end
blanchet@49331
   386
blanchet@49409
   387
fun run_mash_tool ctxt overlord save max_suggs write_cmds read_suggs =
blanchet@49326
   388
  let
blanchet@49409
   389
    val (temp_dir, serial) =
blanchet@49409
   390
      if overlord then (getenv "ISABELLE_HOME_USER", "")
blanchet@49409
   391
      else (getenv "ISABELLE_TMP", serial_string ())
blanchet@49409
   392
    val log_file = if overlord then temp_dir ^ "/mash_log" else "/dev/null"
blanchet@49331
   393
    val err_file = temp_dir ^ "/mash_err" ^ serial
blanchet@49333
   394
    val sugg_file = temp_dir ^ "/mash_suggs" ^ serial
blanchet@49331
   395
    val cmd_file = temp_dir ^ "/mash_commands" ^ serial
blanchet@49409
   396
    val core =
blanchet@49409
   397
      "--inputFile " ^ cmd_file ^ " --predictions " ^ sugg_file ^
blanchet@49409
   398
      " --numberOfPredictions " ^ string_of_int max_suggs ^
blanchet@49409
   399
      (if save then " --saveModel" else "")
blanchet@49409
   400
    val command =
blanchet@49409
   401
      mash_home () ^ "/mash --quiet --outputDir " ^ mash_model_dir () ^
blanchet@49409
   402
      " --log " ^ log_file ^ " " ^ core ^ " >& " ^ err_file
blanchet@49331
   403
  in
blanchet@49338
   404
    write_file ([], K "") sugg_file;
blanchet@49331
   405
    write_file write_cmds cmd_file;
blanchet@49409
   406
    trace_msg ctxt (fn () => "Running " ^ command);
blanchet@49409
   407
    Isabelle_System.bash command;
blanchet@49409
   408
    read_suggs (fn () => try File.read_lines (Path.explode sugg_file) |> these)
blanchet@49409
   409
    |> tap (fn _ => trace_msg ctxt (fn () =>
blanchet@49409
   410
           case try File.read (Path.explode err_file) of
blanchet@49409
   411
             NONE => "Done"
blanchet@49409
   412
           | SOME "" => "Done"
blanchet@49409
   413
           | SOME s => "Error: " ^ elide_string 1000 s))
blanchet@49409
   414
    |> not overlord
blanchet@49409
   415
       ? tap (fn _ => List.app (wipe_out o Path.explode)
blanchet@49409
   416
                               [err_file, sugg_file, cmd_file])
blanchet@49331
   417
  end
blanchet@49331
   418
blanchet@49331
   419
fun str_of_update (name, parents, feats, deps) =
blanchet@49331
   420
  "! " ^ escape_meta name ^ ": " ^ escape_metas parents ^ "; " ^
blanchet@49326
   421
  escape_metas feats ^ "; " ^ escape_metas deps ^ "\n"
blanchet@49326
   422
blanchet@49331
   423
fun str_of_query (parents, feats) =
blanchet@49331
   424
  "? " ^ escape_metas parents ^ "; " ^ escape_metas feats
blanchet@49331
   425
blanchet@49347
   426
fun mash_CLEAR ctxt =
blanchet@49409
   427
  let val path = mash_model_dir () |> Path.explode in
blanchet@49347
   428
    trace_msg ctxt (K "MaSh CLEAR");
blanchet@49409
   429
    File.fold_dir (fn file => fn _ =>
blanchet@49409
   430
                      try File.rm (Path.append path (Path.basic file)))
blanchet@49409
   431
                  path NONE;
blanchet@49409
   432
    ()
blanchet@49324
   433
  end
blanchet@49317
   434
blanchet@49331
   435
fun mash_ADD _ _ [] = ()
blanchet@49331
   436
  | mash_ADD ctxt overlord upds =
blanchet@49331
   437
    (trace_msg ctxt (fn () => "MaSh ADD " ^
blanchet@49331
   438
         elide_string 1000 (space_implode " " (map #1 upds)));
blanchet@49409
   439
     run_mash_tool ctxt overlord true 0 (upds, str_of_update) (K ()))
blanchet@49317
   440
blanchet@49333
   441
fun mash_QUERY ctxt overlord max_suggs (query as (_, feats)) =
blanchet@49329
   442
  (trace_msg ctxt (fn () => "MaSh QUERY " ^ space_implode " " feats);
blanchet@49409
   443
   run_mash_tool ctxt overlord false max_suggs
blanchet@49338
   444
       ([query], str_of_query)
blanchet@49409
   445
       (fn suggs =>
blanchet@49409
   446
           case suggs () of
blanchet@49409
   447
             [] => []
blanchet@49409
   448
           | suggs => snd (extract_query (List.last suggs)))
blanchet@49326
   449
   handle List.Empty => [])
blanchet@49317
   450
blanchet@49317
   451
blanchet@49266
   452
(*** High-level communication with MaSh ***)
blanchet@49266
   453
blanchet@49336
   454
fun try_graph ctxt when def f =
blanchet@49336
   455
  f ()
blanchet@49336
   456
  handle Graph.CYCLES (cycle :: _) =>
blanchet@49336
   457
         (trace_msg ctxt (fn () =>
blanchet@49336
   458
              "Cycle involving " ^ commas cycle ^ " when " ^ when); def)
blanchet@49407
   459
       | Graph.DUP name =>
blanchet@49407
   460
         (trace_msg ctxt (fn () =>
blanchet@49407
   461
              "Duplicate fact " ^ quote name ^ " when " ^ when); def)
blanchet@49336
   462
       | Graph.UNDEF name =>
blanchet@49336
   463
         (trace_msg ctxt (fn () =>
blanchet@49336
   464
              "Unknown fact " ^ quote name ^ " when " ^ when); def)
blanchet@49407
   465
       | exn =>
blanchet@49407
   466
         if Exn.is_interrupt exn then
blanchet@49407
   467
           reraise exn
blanchet@49407
   468
         else
blanchet@49407
   469
           (trace_msg ctxt (fn () =>
blanchet@49407
   470
                "Internal error when " ^ when ^ ":\n" ^
blanchet@49407
   471
                ML_Compiler.exn_message exn); def)
blanchet@49336
   472
blanchet@49405
   473
type mash_state = {fact_graph : unit Graph.T}
blanchet@49264
   474
blanchet@49405
   475
val empty_state = {fact_graph = Graph.empty}
blanchet@49316
   476
blanchet@49316
   477
local
blanchet@49316
   478
blanchet@49405
   479
val version = "*** MaSh 0.0 ***"
blanchet@49405
   480
blanchet@49405
   481
fun load _ (state as (true, _)) = state
blanchet@49405
   482
  | load ctxt _ =
blanchet@49324
   483
    let val path = mash_state_path () in
blanchet@49317
   484
      (true,
blanchet@49317
   485
       case try File.read_lines path of
blanchet@49405
   486
         SOME (version' :: fact_lines) =>
blanchet@49317
   487
         let
blanchet@49337
   488
           fun add_edge_to name parent =
blanchet@49337
   489
             Graph.default_node (parent, ())
blanchet@49337
   490
             #> Graph.add_edge (parent, name)
blanchet@49331
   491
           fun add_fact_line line =
blanchet@49331
   492
             case extract_query line of
blanchet@49331
   493
               ("", _) => I (* shouldn't happen *)
blanchet@49331
   494
             | (name, parents) =>
blanchet@49331
   495
               Graph.default_node (name, ())
blanchet@49337
   496
               #> fold (add_edge_to name) parents
blanchet@49336
   497
           val fact_graph =
blanchet@49337
   498
             try_graph ctxt "loading state" Graph.empty (fn () =>
blanchet@49405
   499
                 Graph.empty |> version' = version
blanchet@49405
   500
                                ? fold add_fact_line fact_lines)
blanchet@49405
   501
         in {fact_graph = fact_graph} end
blanchet@49319
   502
       | _ => empty_state)
blanchet@49317
   503
    end
blanchet@49316
   504
blanchet@49405
   505
fun save {fact_graph} =
blanchet@49316
   506
  let
blanchet@49324
   507
    val path = mash_state_path ()
blanchet@49333
   508
    fun fact_line_for name parents =
blanchet@49333
   509
      escape_meta name ^ ": " ^ escape_metas parents
blanchet@49331
   510
    val append_fact = File.append path o suffix "\n" oo fact_line_for
blanchet@49316
   511
  in
blanchet@49405
   512
    File.write path (version ^ "\n");
blanchet@49331
   513
    Graph.fold (fn (name, ((), (parents, _))) => fn () =>
blanchet@49331
   514
                   append_fact name (Graph.Keys.dest parents))
blanchet@49331
   515
        fact_graph ()
blanchet@49316
   516
  end
blanchet@49316
   517
blanchet@49317
   518
val global_state =
blanchet@49396
   519
  Synchronized.var "Sledgehammer_MaSh.global_state" (false, empty_state)
blanchet@49316
   520
blanchet@49316
   521
in
blanchet@49316
   522
blanchet@49336
   523
fun mash_map ctxt f =
blanchet@49405
   524
  Synchronized.change global_state (load ctxt ##> (f #> tap save))
blanchet@49316
   525
blanchet@49336
   526
fun mash_get ctxt =
blanchet@49405
   527
  Synchronized.change_result global_state (load ctxt #> `snd)
blanchet@49317
   528
blanchet@49347
   529
fun mash_unlearn ctxt =
blanchet@49317
   530
  Synchronized.change global_state (fn _ =>
blanchet@49409
   531
      (mash_CLEAR ctxt; wipe_out (mash_state_path ()); (true, empty_state)))
blanchet@49316
   532
blanchet@49316
   533
end
blanchet@49316
   534
blanchet@49333
   535
fun mash_could_suggest_facts () = mash_home () <> ""
blanchet@49336
   536
fun mash_can_suggest_facts ctxt =
blanchet@49336
   537
  not (Graph.is_empty (#fact_graph (mash_get ctxt)))
blanchet@49264
   538
blanchet@49347
   539
fun parents_wrt_facts facts fact_graph =
blanchet@49331
   540
  let
blanchet@49393
   541
    val facts = [] |> fold (cons o nickname_of o snd) facts
blanchet@49345
   542
    val tab = Symtab.empty |> fold (fn name => Symtab.update (name, ())) facts
blanchet@49392
   543
    fun insert_not_seen seen name =
blanchet@49392
   544
      not (member (op =) seen name) ? insert (op =) name
blanchet@49393
   545
    fun parents_of _ parents [] = parents
blanchet@49392
   546
      | parents_of seen parents (name :: names) =
blanchet@49347
   547
        if Symtab.defined tab name then
blanchet@49392
   548
          parents_of (name :: seen) (name :: parents) names
blanchet@49347
   549
        else
blanchet@49392
   550
          parents_of (name :: seen) parents
blanchet@49392
   551
                     (Graph.Keys.fold (insert_not_seen seen)
blanchet@49392
   552
                                      (Graph.imm_preds fact_graph name) names)
blanchet@49392
   553
  in parents_of [] [] (Graph.maximals fact_graph) end
blanchet@49331
   554
blanchet@49333
   555
(* Generate more suggestions than requested, because some might be thrown out
blanchet@49333
   556
   later for various reasons and "meshing" gives better results with some
blanchet@49333
   557
   slack. *)
blanchet@49333
   558
fun max_suggs_of max_facts = max_facts + Int.min (200, max_facts)
blanchet@49333
   559
blanchet@49335
   560
fun is_fact_in_graph fact_graph (_, th) =
blanchet@49393
   561
  can (Graph.get_node fact_graph) (nickname_of th)
blanchet@49335
   562
blanchet@49333
   563
fun mash_suggest_facts ctxt ({overlord, ...} : params) prover max_facts hyp_ts
blanchet@49333
   564
                       concl_t facts =
blanchet@49316
   565
  let
blanchet@49317
   566
    val thy = Proof_Context.theory_of ctxt
blanchet@49336
   567
    val fact_graph = #fact_graph (mash_get ctxt)
blanchet@49347
   568
    val parents = parents_wrt_facts facts fact_graph
blanchet@49400
   569
    val feats = features_of ctxt prover thy (Local, General) (concl_t :: hyp_ts)
blanchet@49333
   570
    val suggs =
blanchet@49335
   571
      if Graph.is_empty fact_graph then []
blanchet@49335
   572
      else mash_QUERY ctxt overlord (max_suggs_of max_facts) (parents, feats)
blanchet@49335
   573
    val selected = facts |> suggested_facts suggs
blanchet@49335
   574
    val unknown = facts |> filter_out (is_fact_in_graph fact_graph)
blanchet@49335
   575
  in (selected, unknown) end
blanchet@49264
   576
blanchet@49331
   577
fun update_fact_graph ctxt (name, parents, feats, deps) (upds, graph) =
blanchet@49331
   578
  let
blanchet@49331
   579
    fun maybe_add_from from (accum as (parents, graph)) =
blanchet@49336
   580
      try_graph ctxt "updating graph" accum (fn () =>
blanchet@49336
   581
          (from :: parents, Graph.add_edge_acyclic (from, name) graph))
blanchet@49336
   582
    val graph = graph |> Graph.default_node (name, ())
blanchet@49331
   583
    val (parents, graph) = ([], graph) |> fold maybe_add_from parents
blanchet@49331
   584
    val (deps, graph) = ([], graph) |> fold maybe_add_from deps
blanchet@49331
   585
  in ((name, parents, feats, deps) :: upds, graph) end
blanchet@49321
   586
blanchet@49399
   587
val learn_timeout_slack = 2.0
blanchet@49333
   588
blanchet@49399
   589
fun launch_thread timeout task =
blanchet@49398
   590
  let
blanchet@49399
   591
    val hard_timeout = time_mult learn_timeout_slack timeout
blanchet@49399
   592
    val birth_time = Time.now ()
blanchet@49399
   593
    val death_time = Time.+ (birth_time, hard_timeout)
blanchet@49399
   594
    val desc = ("machine learner for Sledgehammer", "")
blanchet@49399
   595
  in Async_Manager.launch MaShN birth_time death_time desc task end
blanchet@49399
   596
blanchet@49399
   597
fun mash_learn_proof ctxt ({overlord, timeout, ...} : params) prover t facts
blanchet@49399
   598
                     used_ths =
blanchet@49399
   599
  if is_smt_prover ctxt prover then
blanchet@49399
   600
    ()
blanchet@49399
   601
  else
blanchet@49399
   602
    launch_thread timeout
blanchet@49399
   603
        (fn () =>
blanchet@49399
   604
            let
blanchet@49399
   605
              val thy = Proof_Context.theory_of ctxt
blanchet@49399
   606
              val name = timestamp () ^ " " ^ serial_string () (* freshish *)
blanchet@49400
   607
              val feats = features_of ctxt prover thy (Local, General) [t]
blanchet@49399
   608
              val deps = used_ths |> map nickname_of
blanchet@49399
   609
            in
blanchet@49405
   610
              mash_map ctxt (fn {fact_graph} =>
blanchet@49399
   611
                  let
blanchet@49399
   612
                    val parents = parents_wrt_facts facts fact_graph
blanchet@49399
   613
                    val upds = [(name, parents, feats, deps)]
blanchet@49399
   614
                    val (upds, fact_graph) =
blanchet@49399
   615
                      ([], fact_graph) |> fold (update_fact_graph ctxt) upds
blanchet@49399
   616
                  in
blanchet@49405
   617
                    mash_ADD ctxt overlord upds; {fact_graph = fact_graph}
blanchet@49399
   618
                  end);
blanchet@49399
   619
              (true, "")
blanchet@49399
   620
            end)
blanchet@49398
   621
blanchet@49407
   622
fun sendback sub =
blanchet@49407
   623
  Markup.markup Isabelle_Markup.sendback (sledgehammerN ^ " " ^ sub)
blanchet@49407
   624
blanchet@49347
   625
(* Too many dependencies is a sign that a decision procedure is at work. There
blanchet@49347
   626
   isn't much too learn from such proofs. *)
blanchet@49347
   627
val max_dependencies = 10
blanchet@49407
   628
val commit_timeout = seconds 30.0
blanchet@49347
   629
blanchet@49333
   630
(* The timeout is understood in a very slack fashion. *)
blanchet@49407
   631
fun mash_learn_facts ctxt (params as {debug, verbose, overlord, timeout, ...})
blanchet@49407
   632
                     prover auto atp learn_timeout facts =
blanchet@49319
   633
  let
blanchet@49333
   634
    val timer = Timer.startRealTimer ()
blanchet@49407
   635
    fun next_commit_time () =
blanchet@49407
   636
      Time.+ (Timer.checkRealTimer timer, commit_timeout)
blanchet@49405
   637
    val {fact_graph} = mash_get ctxt
blanchet@49335
   638
    val new_facts =
blanchet@49335
   639
      facts |> filter_out (is_fact_in_graph fact_graph)
blanchet@49335
   640
            |> sort (thm_ord o pairself snd)
blanchet@49407
   641
    val num_new_facts = length new_facts
blanchet@49323
   642
  in
blanchet@49407
   643
    (if not auto then
blanchet@49409
   644
       "MaShing" ^
blanchet@49409
   645
       (if not auto then
blanchet@49409
   646
          " " ^ string_of_int num_new_facts ^ " fact" ^
blanchet@49409
   647
          plural_s num_new_facts ^
blanchet@49409
   648
          (if atp then " (ATP timeout: " ^ string_from_time timeout ^ ")"
blanchet@49409
   649
           else "")
blanchet@49409
   650
        else
blanchet@49409
   651
          "") ^ "..."
blanchet@49407
   652
     else
blanchet@49409
   653
       "")
blanchet@49407
   654
    |> Output.urgent_message;
blanchet@49323
   655
    if null new_facts then
blanchet@49409
   656
      if not auto then
blanchet@49407
   657
        "Nothing to learn.\n\nHint: Try " ^ sendback relearn_isarN ^ " or " ^
blanchet@49407
   658
        sendback relearn_atpN ^ "  to learn from scratch."
blanchet@49407
   659
      else
blanchet@49407
   660
        ""
blanchet@49323
   661
    else
blanchet@49319
   662
      let
blanchet@49323
   663
        val ths = facts |> map snd
blanchet@49330
   664
        val all_names =
blanchet@49339
   665
          ths |> filter_out is_likely_tautology_or_too_meta
blanchet@49393
   666
              |> map (rpair () o nickname_of)
blanchet@49331
   667
              |> Symtab.make
blanchet@49407
   668
        fun do_commit [] state = state
blanchet@49407
   669
          | do_commit upds {fact_graph} =
blanchet@49407
   670
            let
blanchet@49407
   671
              val (upds, fact_graph) =
blanchet@49407
   672
                ([], fact_graph) |> fold (update_fact_graph ctxt) upds
blanchet@49407
   673
            in mash_ADD ctxt overlord (rev upds); {fact_graph = fact_graph} end
blanchet@49347
   674
        fun trim_deps deps = if length deps > max_dependencies then [] else deps
blanchet@49407
   675
        fun commit last upds =
blanchet@49407
   676
          (if debug andalso not auto then Output.urgent_message "Committing..."
blanchet@49407
   677
           else ();
blanchet@49407
   678
           mash_map ctxt (do_commit (rev upds));
blanchet@49407
   679
           if not last andalso not auto then
blanchet@49407
   680
             let val num_upds = length upds in
blanchet@49407
   681
               "Processed " ^ string_of_int num_upds ^ " fact" ^
blanchet@49407
   682
               plural_s num_upds ^ " in the last " ^
blanchet@49407
   683
               string_from_time commit_timeout ^ "."
blanchet@49407
   684
               |> Output.urgent_message
blanchet@49407
   685
             end
blanchet@49407
   686
           else
blanchet@49407
   687
             ())
blanchet@49407
   688
        fun do_fact _ (accum as (_, (_, _, _, true))) = accum
blanchet@49407
   689
          | do_fact ((_, stature), th)
blanchet@49407
   690
                    (upds, (parents, n, next_commit, false)) =
blanchet@49333
   691
            let
blanchet@49393
   692
              val name = nickname_of th
blanchet@49347
   693
              val feats =
blanchet@49400
   694
                features_of ctxt prover (theory_of_thm th) stature [prop_of th]
blanchet@49407
   695
              val deps =
blanchet@49407
   696
                (if atp then atp_dependencies_of ctxt params prover auto facts
blanchet@49407
   697
                 else isar_dependencies_of) all_names th
blanchet@49407
   698
                |> trim_deps
blanchet@49409
   699
              val n = n |> not (null deps) ? Integer.add 1
blanchet@49407
   700
              val upds = (name, parents, feats, deps) :: upds
blanchet@49407
   701
              val (upds, next_commit) =
blanchet@49407
   702
                if Time.> (Timer.checkRealTimer timer, next_commit) then
blanchet@49407
   703
                  (commit false upds; ([], next_commit_time ()))
blanchet@49407
   704
                else
blanchet@49407
   705
                  (upds, next_commit)
blanchet@49407
   706
              val timed_out =
blanchet@49407
   707
                Time.> (Timer.checkRealTimer timer, learn_timeout)
blanchet@49409
   708
            in (upds, ([name], n, next_commit, timed_out)) end
blanchet@49347
   709
        val parents = parents_wrt_facts facts fact_graph
blanchet@49407
   710
        val (upds, (_, n, _, _)) =
blanchet@49407
   711
          ([], (parents, 0, next_commit_time (), false))
blanchet@49407
   712
          |> fold do_fact new_facts
blanchet@49333
   713
      in
blanchet@49407
   714
        commit true upds;
blanchet@49407
   715
        if verbose orelse not auto then
blanchet@49409
   716
          "Learned " ^ string_of_int n ^ " nontrivial proof" ^ plural_s n ^
blanchet@49334
   717
          (if verbose then
blanchet@49334
   718
             " in " ^ string_from_time (Timer.checkRealTimer timer)
blanchet@49334
   719
           else
blanchet@49334
   720
             "") ^ "."
blanchet@49334
   721
        else
blanchet@49334
   722
          ""
blanchet@49333
   723
      end
blanchet@49323
   724
  end
blanchet@49319
   725
blanchet@49407
   726
fun mash_learn ctxt (params as {provers, ...}) atp =
blanchet@49331
   727
  let
blanchet@49398
   728
    val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
blanchet@49409
   729
    val facts = all_facts_of ctxt css_table
blanchet@49331
   730
  in
blanchet@49407
   731
     mash_learn_facts ctxt params (hd provers) false atp infinite_timeout facts
blanchet@49398
   732
     |> Output.urgent_message
blanchet@49331
   733
  end
blanchet@49264
   734
blanchet@49333
   735
(* The threshold should be large enough so that MaSh doesn't kick in for Auto
blanchet@49333
   736
   Sledgehammer and Try. *)
blanchet@49333
   737
val min_secs_for_learning = 15
blanchet@49333
   738
blanchet@49336
   739
fun relevant_facts ctxt (params as {learn, fact_filter, timeout, ...}) prover
blanchet@49336
   740
        max_facts ({add, only, ...} : fact_override) hyp_ts concl_t facts =
blanchet@49329
   741
  if not (subset (op =) (the_list fact_filter, fact_filters)) then
blanchet@49329
   742
    error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".")
blanchet@49329
   743
  else if only then
blanchet@49304
   744
    facts
blanchet@49336
   745
  else if max_facts <= 0 orelse null facts then
blanchet@49303
   746
    []
blanchet@49303
   747
  else
blanchet@49303
   748
    let
blanchet@49342
   749
      fun maybe_learn () =
blanchet@49399
   750
        if learn andalso not (Async_Manager.has_running_threads MaShN) andalso
blanchet@49399
   751
           Time.toSeconds timeout >= min_secs_for_learning then
blanchet@49399
   752
          let val timeout = time_mult learn_timeout_slack timeout in
blanchet@49399
   753
            launch_thread timeout
blanchet@49407
   754
                (fn () => (true, mash_learn_facts ctxt params prover true false
blanchet@49407
   755
                                                  timeout facts))
blanchet@49334
   756
          end
blanchet@49333
   757
        else
blanchet@49333
   758
          ()
blanchet@49329
   759
      val fact_filter =
blanchet@49329
   760
        case fact_filter of
blanchet@49394
   761
          SOME ff => (() |> ff <> mepoN ? maybe_learn; ff)
blanchet@49333
   762
        | NONE =>
blanchet@49401
   763
          if is_smt_prover ctxt prover then mepoN
blanchet@49401
   764
          else if mash_can_suggest_facts ctxt then (maybe_learn (); meshN)
blanchet@49394
   765
          else if mash_could_suggest_facts () then (maybe_learn (); mepoN)
blanchet@49394
   766
          else mepoN
blanchet@49303
   767
      val add_ths = Attrib.eval_thms ctxt add
blanchet@49307
   768
      fun prepend_facts ths accepts =
blanchet@49303
   769
        ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
blanchet@49307
   770
         (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
blanchet@49308
   771
        |> take max_facts
blanchet@49329
   772
      fun iter () =
blanchet@49313
   773
        iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts
blanchet@49313
   774
                                 concl_t facts
blanchet@49329
   775
      fun mash () =
blanchet@49335
   776
        mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts
blanchet@49329
   777
      val mess =
blanchet@49335
   778
        [] |> (if fact_filter <> mashN then cons (iter (), []) else I)
blanchet@49394
   779
           |> (if fact_filter <> mepoN then cons (mash ()) else I)
blanchet@49303
   780
    in
blanchet@49328
   781
      mesh_facts max_facts mess
blanchet@49303
   782
      |> not (null add_ths) ? prepend_facts add_ths
blanchet@49303
   783
    end
blanchet@49303
   784
blanchet@49334
   785
fun kill_learners () = Async_Manager.kill_threads MaShN "learner"
blanchet@49334
   786
fun running_learners () = Async_Manager.running_threads MaShN "learner"
blanchet@49334
   787
blanchet@49263
   788
end;