src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML
author blanchet
Mon, 23 Jul 2012 15:32:30 +0200
changeset 49456 2d2f009ca8eb
parent 49455 159e320ef851
child 49457 3c9890c19e90
permissions -rw-r--r--
remove MaSh junk associated with size functions
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@49421
    31
  val extract_query : string -> string * (string * real) list
blanchet@49393
    32
  val nickname_of : thm -> string
blanchet@49421
    33
  val suggested_facts :
blanchet@49421
    34
    (string * 'a) list -> ('b * thm) list -> (('b * thm) * 'a) list
blanchet@49336
    35
  val mesh_facts :
blanchet@49421
    36
    int -> ((('a * thm) * real) list * ('a * thm) list) list -> ('a * thm) list
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@49419
    44
  val isar_dependencies_of : unit Symtab.table -> thm -> string list option
blanchet@49407
    45
  val atp_dependencies_of :
blanchet@49419
    46
    Proof.context -> params -> string -> int -> fact list -> unit Symtab.table
blanchet@49419
    47
    -> thm -> string list option
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@49419
    52
  val mash_REPROVE :
blanchet@49419
    53
    Proof.context -> bool -> (string * string list) list -> unit
blanchet@49331
    54
  val mash_QUERY :
blanchet@49421
    55
    Proof.context -> bool -> int -> string list * string list
blanchet@49421
    56
    -> (string * real) list
blanchet@49347
    57
  val mash_unlearn : Proof.context -> unit
blanchet@49333
    58
  val mash_could_suggest_facts : unit -> bool
blanchet@49336
    59
  val mash_can_suggest_facts : Proof.context -> bool
blanchet@49421
    60
  val mash_suggested_facts :
blanchet@49336
    61
    Proof.context -> params -> string -> int -> term list -> term
blanchet@49450
    62
    -> fact list -> (fact * real) list * fact list
blanchet@49398
    63
  val mash_learn_proof :
blanchet@49399
    64
    Proof.context -> params -> string -> term -> ('a * thm) list -> thm list
blanchet@49399
    65
    -> unit
blanchet@49410
    66
  val mash_learn :
blanchet@49410
    67
    Proof.context -> params -> fact_override -> thm list -> bool -> unit
blanchet@49303
    68
  val relevant_facts :
blanchet@49307
    69
    Proof.context -> params -> string -> int -> fact_override -> term list
blanchet@49311
    70
    -> term -> fact list -> fact list
blanchet@49334
    71
  val kill_learners : unit -> unit
blanchet@49334
    72
  val running_learners : unit -> unit
blanchet@49263
    73
end;
blanchet@49263
    74
blanchet@49396
    75
structure Sledgehammer_MaSh : SLEDGEHAMMER_MASH =
blanchet@49263
    76
struct
blanchet@49264
    77
blanchet@49266
    78
open ATP_Util
blanchet@49266
    79
open ATP_Problem_Generate
blanchet@49266
    80
open Sledgehammer_Util
blanchet@49266
    81
open Sledgehammer_Fact
blanchet@49266
    82
open Sledgehammer_Provers
blanchet@49333
    83
open Sledgehammer_Minimize
blanchet@49396
    84
open Sledgehammer_MePo
blanchet@49266
    85
blanchet@49323
    86
val trace =
blanchet@49395
    87
  Attrib.setup_config_bool @{binding sledgehammer_mash_trace} (K false)
blanchet@49323
    88
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
blanchet@49323
    89
blanchet@49334
    90
val MaShN = "MaSh"
blanchet@49334
    91
blanchet@49394
    92
val mepoN = "mepo"
blanchet@49394
    93
val mashN = "mash"
blanchet@49329
    94
val meshN = "mesh"
blanchet@49329
    95
blanchet@49394
    96
val fact_filters = [meshN, mepoN, mashN]
blanchet@49329
    97
blanchet@49407
    98
val unlearnN = "unlearn"
blanchet@49407
    99
val learn_isarN = "learn_isar"
blanchet@49407
   100
val learn_atpN = "learn_atp"
blanchet@49407
   101
val relearn_isarN = "relearn_isar"
blanchet@49407
   102
val relearn_atpN = "relearn_atp"
blanchet@49407
   103
blanchet@49329
   104
fun mash_home () = getenv "MASH_HOME"
blanchet@49409
   105
fun mash_model_dir () =
blanchet@49324
   106
  getenv "ISABELLE_HOME_USER" ^ "/mash"
blanchet@49331
   107
  |> tap (Isabelle_System.mkdir o Path.explode)
blanchet@49409
   108
val mash_state_dir = mash_model_dir
blanchet@49451
   109
fun mash_state_file () = mash_state_dir () ^ "/state"
blanchet@49266
   110
blanchet@49345
   111
blanchet@49266
   112
(*** Isabelle helpers ***)
blanchet@49266
   113
blanchet@49323
   114
fun meta_char c =
blanchet@49266
   115
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
blanchet@49266
   116
     c = #")" orelse c = #"," then
blanchet@49266
   117
    String.str c
blanchet@49266
   118
  else
blanchet@49266
   119
    (* fixed width, in case more digits follow *)
blanchet@49410
   120
    "%" ^ stringN_of_int 3 (Char.ord c)
blanchet@49266
   121
blanchet@49323
   122
fun unmeta_chars accum [] = String.implode (rev accum)
blanchet@49410
   123
  | unmeta_chars accum (#"%" :: d1 :: d2 :: d3 :: cs) =
blanchet@49323
   124
    (case Int.fromString (String.implode [d1, d2, d3]) of
blanchet@49323
   125
       SOME n => unmeta_chars (Char.chr n :: accum) cs
blanchet@49323
   126
     | NONE => "" (* error *))
blanchet@49410
   127
  | unmeta_chars _ (#"%" :: _) = "" (* error *)
blanchet@49323
   128
  | unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs
blanchet@49323
   129
blanchet@49323
   130
val escape_meta = String.translate meta_char
blanchet@49318
   131
val escape_metas = map escape_meta #> space_implode " "
blanchet@49330
   132
val unescape_meta = String.explode #> unmeta_chars []
blanchet@49330
   133
val unescape_metas =
blanchet@49330
   134
  space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta
blanchet@49266
   135
blanchet@49421
   136
fun extract_node line =
blanchet@49421
   137
  case space_explode ":" line of
blanchet@49421
   138
    [name, parents] => (unescape_meta name, unescape_metas parents)
blanchet@49421
   139
  | _ => ("", [])
blanchet@49421
   140
blanchet@49421
   141
fun extract_suggestion sugg =
blanchet@49421
   142
  case space_explode "=" sugg of
blanchet@49421
   143
    [name, weight] =>
blanchet@49421
   144
    SOME (unescape_meta name, Real.fromString weight |> the_default 0.0)
blanchet@49421
   145
  | _ => NONE
blanchet@49421
   146
blanchet@49326
   147
fun extract_query line =
blanchet@49326
   148
  case space_explode ":" line of
blanchet@49421
   149
    [goal, suggs] =>
blanchet@49421
   150
    (unescape_meta goal,
blanchet@49421
   151
     map_filter extract_suggestion (space_explode " " suggs))
blanchet@49327
   152
  | _ => ("", [])
blanchet@49326
   153
blanchet@49393
   154
fun parent_of_local_thm th =
blanchet@49393
   155
  let
blanchet@49393
   156
    val thy = th |> Thm.theory_of_thm
blanchet@49393
   157
    val facts = thy |> Global_Theory.facts_of
blanchet@49393
   158
    val space = facts |> Facts.space_of
blanchet@49393
   159
    fun id_of s = #id (Name_Space.the_entry space s)
blanchet@49393
   160
    fun max_id (s', _) (s, id) =
blanchet@49393
   161
      let val id' = id_of s' in if id > id' then (s, id) else (s', id') end
blanchet@49393
   162
  in ("", ~1) |> Facts.fold_static max_id facts |> fst end
blanchet@49393
   163
blanchet@49393
   164
val local_prefix = "local" ^ Long_Name.separator
blanchet@49393
   165
blanchet@49393
   166
fun nickname_of th =
blanchet@49409
   167
  if Thm.has_name_hint th then
blanchet@49409
   168
    let val hint = Thm.get_name_hint th in
blanchet@49409
   169
      (* FIXME: There must be a better way to detect local facts. *)
blanchet@49409
   170
      case try (unprefix local_prefix) hint of
blanchet@49409
   171
        SOME suf =>
blanchet@49409
   172
        parent_of_local_thm th ^ Long_Name.separator ^ Long_Name.separator ^ suf
blanchet@49409
   173
      | NONE => hint
blanchet@49409
   174
    end
blanchet@49409
   175
  else
blanchet@49409
   176
    backquote_thm th
blanchet@49393
   177
blanchet@49345
   178
fun suggested_facts suggs facts =
blanchet@49345
   179
  let
blanchet@49393
   180
    fun add_fact (fact as (_, th)) = Symtab.default (nickname_of th, fact)
blanchet@49345
   181
    val tab = Symtab.empty |> fold add_fact facts
blanchet@49421
   182
    fun find_sugg (name, weight) =
blanchet@49421
   183
      Symtab.lookup tab name |> Option.map (rpair weight)
blanchet@49421
   184
  in map_filter find_sugg suggs end
blanchet@49326
   185
blanchet@49421
   186
fun sum_avg [] = 0
blanchet@49422
   187
  | sum_avg xs =
blanchet@49422
   188
    Real.ceil (100000000.0 * fold (curry (op +)) xs 0.0) div length xs
blanchet@49343
   189
blanchet@49421
   190
fun normalize_scores [] = []
blanchet@49421
   191
  | normalize_scores ((fact, score) :: tail) =
blanchet@49421
   192
    (fact, 1.0) :: map (apsnd (curry Real.* (1.0 / score))) tail
blanchet@49328
   193
blanchet@49421
   194
fun mesh_facts max_facts [(sels, unks)] =
blanchet@49421
   195
    map fst (take max_facts sels) @ take (max_facts - length sels) unks
blanchet@49329
   196
  | mesh_facts max_facts mess =
blanchet@49329
   197
    let
blanchet@49421
   198
      val mess = mess |> map (apfst (normalize_scores #> `length))
blanchet@49329
   199
      val fact_eq = Thm.eq_thm o pairself snd
blanchet@49421
   200
      fun score_at sels = try (nth sels) #> Option.map snd
blanchet@49335
   201
      fun score_in fact ((sel_len, sels), unks) =
blanchet@49421
   202
        case find_index (curry fact_eq fact o fst) sels of
blanchet@49335
   203
          ~1 => (case find_index (curry fact_eq fact) unks of
blanchet@49421
   204
                   ~1 => score_at sels sel_len
blanchet@49335
   205
                 | _ => NONE)
blanchet@49421
   206
        | rank => score_at sels rank
blanchet@49421
   207
      fun weight_of fact = mess |> map_filter (score_in fact) |> sum_avg
blanchet@49421
   208
      val facts =
blanchet@49421
   209
        fold (union fact_eq o map fst o take max_facts o snd o fst) mess []
blanchet@49329
   210
    in
blanchet@49421
   211
      facts |> map (`weight_of) |> sort (int_ord o swap o pairself fst)
blanchet@49343
   212
            |> map snd |> take max_facts
blanchet@49329
   213
    end
blanchet@49327
   214
blanchet@49410
   215
val thy_feature_name_of = prefix "y"
blanchet@49410
   216
val const_name_of = prefix "c"
blanchet@49410
   217
val type_name_of = prefix "t"
blanchet@49410
   218
val class_name_of = prefix "s"
blanchet@49266
   219
blanchet@49339
   220
fun theory_ord p =
blanchet@49339
   221
  if Theory.eq_thy p then
blanchet@49339
   222
    EQUAL
blanchet@49339
   223
  else if Theory.subthy p then
blanchet@49339
   224
    LESS
blanchet@49339
   225
  else if Theory.subthy (swap p) then
blanchet@49339
   226
    GREATER
blanchet@49339
   227
  else case int_ord (pairself (length o Theory.ancestors_of) p) of
blanchet@49339
   228
    EQUAL => string_ord (pairself Context.theory_name p)
blanchet@49339
   229
  | order => order
blanchet@49339
   230
blanchet@49339
   231
val thm_ord = theory_ord o pairself theory_of_thm
blanchet@49339
   232
blanchet@49407
   233
val freezeT = Type.legacy_freeze_type
blanchet@49407
   234
blanchet@49407
   235
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49407
   236
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49407
   237
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49407
   238
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49407
   239
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49407
   240
  | freeze t = t
blanchet@49407
   241
blanchet@49407
   242
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
blanchet@49407
   243
blanchet@49407
   244
fun run_prover_for_mash ctxt params prover facts goal =
blanchet@49407
   245
  let
blanchet@49407
   246
    val problem =
blanchet@49407
   247
      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
blanchet@49407
   248
       facts = facts |> map (apfst (apfst (fn name => name ())))
blanchet@49407
   249
                     |> map Untranslated_Fact}
blanchet@49407
   250
  in
blanchet@49414
   251
    get_minimizing_prover ctxt MaSh (K (K ())) prover params (K (K (K "")))
blanchet@49407
   252
                          problem
blanchet@49407
   253
  end
blanchet@49407
   254
blanchet@49341
   255
val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49341
   256
blanchet@49413
   257
val logical_consts =
blanchet@49413
   258
  [@{const_name prop}, @{const_name Pure.conjunction}] @ atp_logical_consts
blanchet@49413
   259
blanchet@49333
   260
fun interesting_terms_types_and_classes ctxt prover term_max_depth
blanchet@49333
   261
                                        type_max_depth ts =
blanchet@49266
   262
  let
blanchet@49333
   263
    fun is_bad_const (x as (s, _)) args =
blanchet@49413
   264
      member (op =) logical_consts s orelse
blanchet@49333
   265
      fst (is_built_in_const_for_prover ctxt prover x args)
blanchet@49319
   266
    fun add_classes @{sort type} = I
blanchet@49319
   267
      | add_classes S = union (op =) (map class_name_of S)
blanchet@49266
   268
    fun do_add_type (Type (s, Ts)) =
blanchet@49266
   269
        (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
blanchet@49266
   270
        #> fold do_add_type Ts
blanchet@49319
   271
      | do_add_type (TFree (_, S)) = add_classes S
blanchet@49319
   272
      | do_add_type (TVar (_, S)) = add_classes S
blanchet@49266
   273
    fun add_type T = type_max_depth >= 0 ? do_add_type T
blanchet@49266
   274
    fun mk_app s args =
blanchet@49266
   275
      if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
blanchet@49266
   276
      else s
blanchet@49266
   277
    fun patternify ~1 _ = ""
blanchet@49266
   278
      | patternify depth t =
blanchet@49266
   279
        case strip_comb t of
blanchet@49413
   280
          (Const (x as (s, _)), args) =>
blanchet@49413
   281
          if is_bad_const x args then ""
blanchet@49413
   282
          else mk_app (const_name_of s) (map (patternify (depth - 1)) args)
blanchet@49266
   283
        | _ => ""
blanchet@49421
   284
    fun add_pattern depth t =
blanchet@49421
   285
      case patternify depth t of "" => I | s => insert (op =) s
blanchet@49266
   286
    fun add_term_patterns ~1 _ = I
blanchet@49266
   287
      | add_term_patterns depth t =
blanchet@49421
   288
        add_pattern depth t #> add_term_patterns (depth - 1) t
blanchet@49266
   289
    val add_term = add_term_patterns term_max_depth
blanchet@49266
   290
    fun add_patterns t =
blanchet@49266
   291
      let val (head, args) = strip_comb t in
blanchet@49266
   292
        (case head of
blanchet@49413
   293
           Const (_, T) => add_term t #> add_type T
blanchet@49266
   294
         | Free (_, T) => add_type T
blanchet@49266
   295
         | Var (_, T) => add_type T
blanchet@49266
   296
         | Abs (_, T, body) => add_type T #> add_patterns body
blanchet@49266
   297
         | _ => I)
blanchet@49266
   298
        #> fold add_patterns args
blanchet@49266
   299
      end
blanchet@49341
   300
  in [] |> fold add_patterns ts end
blanchet@49266
   301
blanchet@49266
   302
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
blanchet@49266
   303
blanchet@49312
   304
val term_max_depth = 1
blanchet@49312
   305
val type_max_depth = 1
blanchet@49266
   306
blanchet@49266
   307
(* TODO: Generate type classes for types? *)
blanchet@49400
   308
fun features_of ctxt prover thy (scope, status) ts =
blanchet@49318
   309
  thy_feature_name_of (Context.theory_name thy) ::
blanchet@49333
   310
  interesting_terms_types_and_classes ctxt prover term_max_depth type_max_depth
blanchet@49333
   311
                                      ts
blanchet@49347
   312
  |> forall is_lambda_free ts ? cons "no_lams"
blanchet@49347
   313
  |> forall (not o exists_Const is_exists) ts ? cons "no_skos"
blanchet@49400
   314
  |> scope <> Global ? cons "local"
blanchet@49453
   315
  |> (case string_of_status status of "" => I | s => cons s)
blanchet@49266
   316
blanchet@49419
   317
(* Too many dependencies is a sign that a decision procedure is at work. There
blanchet@49419
   318
   isn't much too learn from such proofs. *)
blanchet@49455
   319
val max_dependencies = 20
blanchet@49419
   320
val atp_dependency_default_max_fact = 50
blanchet@49266
   321
blanchet@49453
   322
(* "type_definition_xxx" facts are characterized by their use of "CollectI". *)
blanchet@49456
   323
val typedef_deps = [@{thm CollectI} |> nickname_of]
blanchet@49453
   324
blanchet@49453
   325
(* "Rep_xxx_inject", "Abs_xxx_inverse", etc., are derived using these facts. *)
blanchet@49453
   326
val typedef_ths =
blanchet@49453
   327
  @{thms type_definition.Abs_inverse type_definition.Rep_inverse
blanchet@49453
   328
         type_definition.Rep type_definition.Rep_inject
blanchet@49453
   329
         type_definition.Abs_inject type_definition.Rep_cases
blanchet@49453
   330
         type_definition.Abs_cases type_definition.Rep_induct
blanchet@49453
   331
         type_definition.Abs_induct type_definition.Rep_range
blanchet@49453
   332
         type_definition.Abs_image}
blanchet@49453
   333
  |> map nickname_of
blanchet@49453
   334
blanchet@49456
   335
fun is_size_def [dep] th =
blanchet@49456
   336
    (case first_field ".recs" dep of
blanchet@49456
   337
       SOME (pref, _) =>
blanchet@49456
   338
       (case first_field ".size" (nickname_of th) of
blanchet@49456
   339
          SOME (pref', _) => pref = pref'
blanchet@49456
   340
        | NONE => false)
blanchet@49456
   341
     | NONE => false)
blanchet@49456
   342
  | is_size_def _ _ = false
blanchet@49456
   343
blanchet@49456
   344
fun trim_dependencies th deps =
blanchet@49456
   345
  if length deps > max_dependencies orelse deps = typedef_deps orelse
blanchet@49456
   346
     exists (member (op =) typedef_ths) deps orelse is_size_def deps th then
blanchet@49453
   347
    NONE
blanchet@49453
   348
  else
blanchet@49453
   349
    SOME deps
blanchet@49266
   350
blanchet@49456
   351
fun isar_dependencies_of all_names th =
blanchet@49456
   352
  th |> thms_in_proof (SOME all_names) |> trim_dependencies th
blanchet@49419
   353
blanchet@49419
   354
fun atp_dependencies_of ctxt (params as {verbose, max_facts, ...}) prover
blanchet@49419
   355
                        auto_level facts all_names th =
blanchet@49407
   356
  case isar_dependencies_of all_names th of
blanchet@49419
   357
    SOME [] => NONE
blanchet@49407
   358
  | isar_deps =>
blanchet@49407
   359
    let
blanchet@49407
   360
      val thy = Proof_Context.theory_of ctxt
blanchet@49407
   361
      val goal = goal_of_thm thy th
blanchet@49407
   362
      val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
blanchet@49407
   363
      val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
blanchet@49407
   364
      fun fix_name ((_, stature), th) = ((fn () => nickname_of th, stature), th)
blanchet@49407
   365
      fun is_dep dep (_, th) = nickname_of th = dep
blanchet@49407
   366
      fun add_isar_dep facts dep accum =
blanchet@49407
   367
        if exists (is_dep dep) accum then
blanchet@49407
   368
          accum
blanchet@49407
   369
        else case find_first (is_dep dep) facts of
blanchet@49407
   370
          SOME ((name, status), th) => accum @ [((name, status), th)]
blanchet@49407
   371
        | NONE => accum (* shouldn't happen *)
blanchet@49407
   372
      val facts =
blanchet@49421
   373
        facts |> mepo_suggested_facts ctxt params prover
blanchet@49419
   374
                     (max_facts |> the_default atp_dependency_default_max_fact)
blanchet@49419
   375
                     NONE hyp_ts concl_t
blanchet@49419
   376
              |> fold (add_isar_dep facts) (these isar_deps)
blanchet@49407
   377
              |> map fix_name
blanchet@49407
   378
    in
blanchet@49419
   379
      if verbose andalso auto_level = 0 then
blanchet@49407
   380
        let val num_facts = length facts in
blanchet@49407
   381
          "MaSh: " ^ quote prover ^ " on " ^ quote (nickname_of th) ^
blanchet@49407
   382
          " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
blanchet@49407
   383
          "."
blanchet@49407
   384
          |> Output.urgent_message
blanchet@49407
   385
        end
blanchet@49407
   386
      else
blanchet@49407
   387
        ();
blanchet@49407
   388
      case run_prover_for_mash ctxt params prover facts goal of
blanchet@49407
   389
        {outcome = NONE, used_facts, ...} =>
blanchet@49419
   390
        (if verbose andalso auto_level = 0 then
blanchet@49407
   391
           let val num_facts = length used_facts in
blanchet@49407
   392
             "Found proof with " ^ string_of_int num_facts ^ " fact" ^
blanchet@49407
   393
             plural_s num_facts ^ "."
blanchet@49407
   394
             |> Output.urgent_message
blanchet@49407
   395
           end
blanchet@49407
   396
         else
blanchet@49407
   397
           ();
blanchet@49456
   398
         used_facts |> map fst |> trim_dependencies th)
blanchet@49419
   399
      | _ => NONE
blanchet@49407
   400
    end
blanchet@49266
   401
blanchet@49266
   402
blanchet@49317
   403
(*** Low-level communication with MaSh ***)
blanchet@49317
   404
blanchet@49409
   405
(* more friendly than "try o File.rm" for those who keep the files open in their
blanchet@49409
   406
   text editor *)
blanchet@49451
   407
fun wipe_out_file file = File.write (Path.explode file) ""
blanchet@49409
   408
blanchet@49451
   409
fun write_file heading (xs, f) file =
blanchet@49333
   410
  let val path = Path.explode file in
blanchet@49451
   411
    File.write path heading;
blanchet@49338
   412
    xs |> chunk_list 500
blanchet@49338
   413
       |> List.app (File.append path o space_implode "" o map f)
blanchet@49333
   414
  end
blanchet@49331
   415
blanchet@49409
   416
fun run_mash_tool ctxt overlord save max_suggs write_cmds read_suggs =
blanchet@49326
   417
  let
blanchet@49409
   418
    val (temp_dir, serial) =
blanchet@49409
   419
      if overlord then (getenv "ISABELLE_HOME_USER", "")
blanchet@49409
   420
      else (getenv "ISABELLE_TMP", serial_string ())
blanchet@49409
   421
    val log_file = if overlord then temp_dir ^ "/mash_log" else "/dev/null"
blanchet@49331
   422
    val err_file = temp_dir ^ "/mash_err" ^ serial
blanchet@49333
   423
    val sugg_file = temp_dir ^ "/mash_suggs" ^ serial
blanchet@49331
   424
    val cmd_file = temp_dir ^ "/mash_commands" ^ serial
blanchet@49409
   425
    val core =
blanchet@49409
   426
      "--inputFile " ^ cmd_file ^ " --predictions " ^ sugg_file ^
blanchet@49409
   427
      " --numberOfPredictions " ^ string_of_int max_suggs ^
blanchet@49409
   428
      (if save then " --saveModel" else "")
blanchet@49409
   429
    val command =
blanchet@49409
   430
      mash_home () ^ "/mash --quiet --outputDir " ^ mash_model_dir () ^
blanchet@49409
   431
      " --log " ^ log_file ^ " " ^ core ^ " >& " ^ err_file
blanchet@49331
   432
  in
blanchet@49451
   433
    write_file "" ([], K "") sugg_file;
blanchet@49451
   434
    write_file "" write_cmds cmd_file;
blanchet@49409
   435
    trace_msg ctxt (fn () => "Running " ^ command);
blanchet@49409
   436
    Isabelle_System.bash command;
blanchet@49409
   437
    read_suggs (fn () => try File.read_lines (Path.explode sugg_file) |> these)
blanchet@49409
   438
    |> tap (fn _ => trace_msg ctxt (fn () =>
blanchet@49409
   439
           case try File.read (Path.explode err_file) of
blanchet@49409
   440
             NONE => "Done"
blanchet@49409
   441
           | SOME "" => "Done"
blanchet@49409
   442
           | SOME s => "Error: " ^ elide_string 1000 s))
blanchet@49409
   443
    |> not overlord
blanchet@49451
   444
       ? tap (fn _ => List.app wipe_out_file [err_file, sugg_file, cmd_file])
blanchet@49331
   445
  end
blanchet@49331
   446
blanchet@49419
   447
fun str_of_add (name, parents, feats, deps) =
blanchet@49331
   448
  "! " ^ escape_meta name ^ ": " ^ escape_metas parents ^ "; " ^
blanchet@49326
   449
  escape_metas feats ^ "; " ^ escape_metas deps ^ "\n"
blanchet@49326
   450
blanchet@49419
   451
fun str_of_reprove (name, deps) =
blanchet@49419
   452
  "p " ^ escape_meta name ^ ": " ^ escape_metas deps ^ "\n"
blanchet@49419
   453
blanchet@49331
   454
fun str_of_query (parents, feats) =
blanchet@49421
   455
  "? " ^ escape_metas parents ^ "; " ^ escape_metas feats ^ "\n"
blanchet@49331
   456
blanchet@49347
   457
fun mash_CLEAR ctxt =
blanchet@49409
   458
  let val path = mash_model_dir () |> Path.explode in
blanchet@49347
   459
    trace_msg ctxt (K "MaSh CLEAR");
blanchet@49409
   460
    File.fold_dir (fn file => fn _ =>
blanchet@49409
   461
                      try File.rm (Path.append path (Path.basic file)))
blanchet@49409
   462
                  path NONE;
blanchet@49409
   463
    ()
blanchet@49324
   464
  end
blanchet@49317
   465
blanchet@49331
   466
fun mash_ADD _ _ [] = ()
blanchet@49419
   467
  | mash_ADD ctxt overlord adds =
blanchet@49331
   468
    (trace_msg ctxt (fn () => "MaSh ADD " ^
blanchet@49419
   469
         elide_string 1000 (space_implode " " (map #1 adds)));
blanchet@49419
   470
     run_mash_tool ctxt overlord true 0 (adds, str_of_add) (K ()))
blanchet@49419
   471
blanchet@49419
   472
fun mash_REPROVE _ _ [] = ()
blanchet@49419
   473
  | mash_REPROVE ctxt overlord reps =
blanchet@49419
   474
    (trace_msg ctxt (fn () => "MaSh REPROVE " ^
blanchet@49419
   475
         elide_string 1000 (space_implode " " (map #1 reps)));
blanchet@49419
   476
     run_mash_tool ctxt overlord true 0 (reps, str_of_reprove) (K ()))
blanchet@49317
   477
blanchet@49333
   478
fun mash_QUERY ctxt overlord max_suggs (query as (_, feats)) =
blanchet@49329
   479
  (trace_msg ctxt (fn () => "MaSh QUERY " ^ space_implode " " feats);
blanchet@49409
   480
   run_mash_tool ctxt overlord false max_suggs
blanchet@49338
   481
       ([query], str_of_query)
blanchet@49409
   482
       (fn suggs =>
blanchet@49409
   483
           case suggs () of
blanchet@49409
   484
             [] => []
blanchet@49409
   485
           | suggs => snd (extract_query (List.last suggs)))
blanchet@49326
   486
   handle List.Empty => [])
blanchet@49317
   487
blanchet@49317
   488
blanchet@49266
   489
(*** High-level communication with MaSh ***)
blanchet@49266
   490
blanchet@49336
   491
fun try_graph ctxt when def f =
blanchet@49336
   492
  f ()
blanchet@49336
   493
  handle Graph.CYCLES (cycle :: _) =>
blanchet@49336
   494
         (trace_msg ctxt (fn () =>
blanchet@49336
   495
              "Cycle involving " ^ commas cycle ^ " when " ^ when); def)
blanchet@49407
   496
       | Graph.DUP name =>
blanchet@49407
   497
         (trace_msg ctxt (fn () =>
blanchet@49407
   498
              "Duplicate fact " ^ quote name ^ " when " ^ when); def)
blanchet@49336
   499
       | Graph.UNDEF name =>
blanchet@49336
   500
         (trace_msg ctxt (fn () =>
blanchet@49336
   501
              "Unknown fact " ^ quote name ^ " when " ^ when); def)
blanchet@49407
   502
       | exn =>
blanchet@49407
   503
         if Exn.is_interrupt exn then
blanchet@49407
   504
           reraise exn
blanchet@49407
   505
         else
blanchet@49407
   506
           (trace_msg ctxt (fn () =>
blanchet@49407
   507
                "Internal error when " ^ when ^ ":\n" ^
blanchet@49407
   508
                ML_Compiler.exn_message exn); def)
blanchet@49336
   509
blanchet@49421
   510
fun graph_info G =
blanchet@49421
   511
  string_of_int (length (Graph.keys G)) ^ " node(s), " ^
blanchet@49421
   512
  string_of_int (fold (Integer.add o length o snd) (Graph.dest G) 0) ^
blanchet@49421
   513
  " edge(s), " ^
blanchet@49421
   514
  string_of_int (length (Graph.minimals G)) ^ " minimal, " ^
blanchet@49421
   515
  string_of_int (length (Graph.maximals G)) ^ " maximal"
blanchet@49421
   516
blanchet@49415
   517
type mash_state = {fact_G : unit Graph.T}
blanchet@49264
   518
blanchet@49415
   519
val empty_state = {fact_G = Graph.empty}
blanchet@49316
   520
blanchet@49316
   521
local
blanchet@49316
   522
blanchet@49405
   523
val version = "*** MaSh 0.0 ***"
blanchet@49405
   524
blanchet@49405
   525
fun load _ (state as (true, _)) = state
blanchet@49405
   526
  | load ctxt _ =
blanchet@49451
   527
    let val path = mash_state_file () |> Path.explode in
blanchet@49317
   528
      (true,
blanchet@49317
   529
       case try File.read_lines path of
blanchet@49421
   530
         SOME (version' :: node_lines) =>
blanchet@49317
   531
         let
blanchet@49337
   532
           fun add_edge_to name parent =
blanchet@49421
   533
             Graph.default_node (parent, ()) #> Graph.add_edge (parent, name)
blanchet@49421
   534
           fun add_node line =
blanchet@49421
   535
             case extract_node line of
blanchet@49331
   536
               ("", _) => I (* shouldn't happen *)
blanchet@49331
   537
             | (name, parents) =>
blanchet@49421
   538
               Graph.default_node (name, ()) #> fold (add_edge_to name) parents
blanchet@49415
   539
           val fact_G =
blanchet@49337
   540
             try_graph ctxt "loading state" Graph.empty (fn () =>
blanchet@49421
   541
                 Graph.empty |> version' = version ? fold add_node node_lines)
blanchet@49421
   542
         in
blanchet@49421
   543
           trace_msg ctxt (fn () =>
blanchet@49421
   544
               "Loaded fact graph (" ^ graph_info fact_G ^ ")");
blanchet@49421
   545
           {fact_G = fact_G}
blanchet@49421
   546
         end
blanchet@49319
   547
       | _ => empty_state)
blanchet@49317
   548
    end
blanchet@49316
   549
blanchet@49421
   550
fun save ctxt {fact_G} =
blanchet@49316
   551
  let
blanchet@49451
   552
    fun str_of_entry (name, parents) =
blanchet@49451
   553
      escape_meta name ^ ": " ^ escape_metas parents ^ "\n"
blanchet@49451
   554
    fun append_entry (name, ((), (parents, _))) =
blanchet@49451
   555
      cons (name, Graph.Keys.dest parents)
blanchet@49451
   556
    val entries = [] |> Graph.fold append_entry fact_G
blanchet@49316
   557
  in
blanchet@49451
   558
    write_file (version ^ "\n") (entries, str_of_entry) (mash_state_file ());
blanchet@49421
   559
    trace_msg ctxt (fn () => "Saved fact graph (" ^ graph_info fact_G ^ ")")
blanchet@49316
   560
  end
blanchet@49316
   561
blanchet@49317
   562
val global_state =
blanchet@49396
   563
  Synchronized.var "Sledgehammer_MaSh.global_state" (false, empty_state)
blanchet@49316
   564
blanchet@49316
   565
in
blanchet@49316
   566
blanchet@49336
   567
fun mash_map ctxt f =
blanchet@49421
   568
  Synchronized.change global_state (load ctxt ##> (f #> tap (save ctxt)))
blanchet@49316
   569
blanchet@49449
   570
fun mash_peek ctxt f =
blanchet@49449
   571
  Synchronized.change_result global_state (load ctxt #> `snd #>> f)
blanchet@49449
   572
blanchet@49336
   573
fun mash_get ctxt =
blanchet@49405
   574
  Synchronized.change_result global_state (load ctxt #> `snd)
blanchet@49317
   575
blanchet@49347
   576
fun mash_unlearn ctxt =
blanchet@49317
   577
  Synchronized.change global_state (fn _ =>
blanchet@49451
   578
      (mash_CLEAR ctxt;
blanchet@49451
   579
       wipe_out_file (mash_state_file ());
blanchet@49451
   580
       (true, empty_state)))
blanchet@49316
   581
blanchet@49316
   582
end
blanchet@49316
   583
blanchet@49333
   584
fun mash_could_suggest_facts () = mash_home () <> ""
blanchet@49415
   585
fun mash_can_suggest_facts ctxt = not (Graph.is_empty (#fact_G (mash_get ctxt)))
blanchet@49264
   586
blanchet@49422
   587
fun num_keys keys = Graph.Keys.fold (K (Integer.add 1)) keys 0
blanchet@49415
   588
blanchet@49422
   589
fun maximal_in_graph fact_G facts =
blanchet@49331
   590
  let
blanchet@49393
   591
    val facts = [] |> fold (cons o nickname_of o snd) facts
blanchet@49422
   592
    val tab = Symtab.empty |> fold (fn name => Symtab.default (name, ())) facts
blanchet@49422
   593
    fun insert_new seen name =
blanchet@49422
   594
      not (Symtab.defined seen name) ? insert (op =) name
blanchet@49422
   595
    fun find_maxes _ (maxs, []) = map snd maxs
blanchet@49422
   596
      | find_maxes seen (maxs, new :: news) =
blanchet@49422
   597
        find_maxes
blanchet@49422
   598
            (seen |> num_keys (Graph.imm_succs fact_G new) > 1
blanchet@49422
   599
                     ? Symtab.default (new, ()))
blanchet@49422
   600
            (if Symtab.defined tab new then
blanchet@49422
   601
               let
blanchet@49422
   602
                 val newp = Graph.all_preds fact_G [new]
blanchet@49422
   603
                 fun is_ancestor x yp = member (op =) yp x
blanchet@49422
   604
                 val maxs =
blanchet@49422
   605
                   maxs |> filter (fn (_, max) => not (is_ancestor max newp))
blanchet@49422
   606
               in
blanchet@49422
   607
                 if exists (is_ancestor new o fst) maxs then
blanchet@49422
   608
                   (maxs, news)
blanchet@49422
   609
                 else
blanchet@49422
   610
                   ((newp, new)
blanchet@49422
   611
                    :: filter_out (fn (_, max) => is_ancestor max newp) maxs,
blanchet@49422
   612
                    news)
blanchet@49422
   613
               end
blanchet@49422
   614
             else
blanchet@49422
   615
               (maxs, Graph.Keys.fold (insert_new seen)
blanchet@49422
   616
                                      (Graph.imm_preds fact_G new) news))
blanchet@49422
   617
  in find_maxes Symtab.empty ([], Graph.maximals fact_G) end
blanchet@49331
   618
blanchet@49333
   619
(* Generate more suggestions than requested, because some might be thrown out
blanchet@49333
   620
   later for various reasons and "meshing" gives better results with some
blanchet@49333
   621
   slack. *)
blanchet@49449
   622
fun max_suggs_of max_facts = max_facts + Int.min (50, max_facts)
blanchet@49333
   623
blanchet@49415
   624
fun is_fact_in_graph fact_G (_, th) =
blanchet@49415
   625
  can (Graph.get_node fact_G) (nickname_of th)
blanchet@49335
   626
blanchet@49450
   627
fun interleave [] ys = ys
blanchet@49450
   628
  | interleave xs [] = xs
blanchet@49450
   629
  | interleave (x :: xs) (y :: ys) = x :: y :: interleave xs ys
blanchet@49450
   630
blanchet@49421
   631
fun mash_suggested_facts ctxt ({overlord, ...} : params) prover max_facts hyp_ts
blanchet@49421
   632
                         concl_t facts =
blanchet@49316
   633
  let
blanchet@49317
   634
    val thy = Proof_Context.theory_of ctxt
blanchet@49449
   635
    val (fact_G, suggs) =
blanchet@49449
   636
      mash_peek ctxt (fn {fact_G} =>
blanchet@49449
   637
          if Graph.is_empty fact_G then
blanchet@49449
   638
            (fact_G, [])
blanchet@49449
   639
          else
blanchet@49449
   640
            let
blanchet@49449
   641
              val parents = maximal_in_graph fact_G facts
blanchet@49449
   642
              val feats =
blanchet@49449
   643
                features_of ctxt prover thy (Local, General) (concl_t :: hyp_ts)
blanchet@49449
   644
            in
blanchet@49449
   645
              (fact_G, mash_QUERY ctxt overlord (max_suggs_of max_facts)
blanchet@49449
   646
                                  (parents, feats))
blanchet@49449
   647
            end)
blanchet@49450
   648
    val sels =
blanchet@49423
   649
      facts |> suggested_facts suggs
blanchet@49423
   650
            (* The weights currently returned by "mash.py" are too extreme to
blanchet@49423
   651
               make any sense. *)
blanchet@49450
   652
            |> map fst
blanchet@49450
   653
    val (unk_global, unk_local) =
blanchet@49450
   654
      facts |> filter_out (is_fact_in_graph fact_G)
blanchet@49450
   655
            |> List.partition (fn ((_, (loc, _)), _) => loc = Global)
blanchet@49450
   656
  in (interleave unk_local sels |> weight_mepo_facts, unk_global) end
blanchet@49264
   657
blanchet@49419
   658
fun add_to_fact_graph ctxt (name, parents, feats, deps) (adds, graph) =
blanchet@49331
   659
  let
blanchet@49331
   660
    fun maybe_add_from from (accum as (parents, graph)) =
blanchet@49336
   661
      try_graph ctxt "updating graph" accum (fn () =>
blanchet@49336
   662
          (from :: parents, Graph.add_edge_acyclic (from, name) graph))
blanchet@49336
   663
    val graph = graph |> Graph.default_node (name, ())
blanchet@49331
   664
    val (parents, graph) = ([], graph) |> fold maybe_add_from parents
blanchet@49422
   665
    val (deps, _) = ([], graph) |> fold maybe_add_from deps
blanchet@49419
   666
  in ((name, parents, feats, deps) :: adds, graph) end
blanchet@49321
   667
blanchet@49399
   668
val learn_timeout_slack = 2.0
blanchet@49333
   669
blanchet@49399
   670
fun launch_thread timeout task =
blanchet@49398
   671
  let
blanchet@49399
   672
    val hard_timeout = time_mult learn_timeout_slack timeout
blanchet@49399
   673
    val birth_time = Time.now ()
blanchet@49399
   674
    val death_time = Time.+ (birth_time, hard_timeout)
blanchet@49399
   675
    val desc = ("machine learner for Sledgehammer", "")
blanchet@49399
   676
  in Async_Manager.launch MaShN birth_time death_time desc task end
blanchet@49399
   677
blanchet@49415
   678
fun freshish_name () =
blanchet@49415
   679
  Date.fmt ".%Y_%m_%d_%H_%M_%S__" (Date.fromTimeLocal (Time.now ())) ^
blanchet@49415
   680
  serial_string ()
blanchet@49415
   681
blanchet@49399
   682
fun mash_learn_proof ctxt ({overlord, timeout, ...} : params) prover t facts
blanchet@49399
   683
                     used_ths =
blanchet@49399
   684
  if is_smt_prover ctxt prover then
blanchet@49399
   685
    ()
blanchet@49399
   686
  else
blanchet@49418
   687
    launch_thread timeout (fn () =>
blanchet@49418
   688
        let
blanchet@49418
   689
          val thy = Proof_Context.theory_of ctxt
blanchet@49418
   690
          val name = freshish_name ()
blanchet@49418
   691
          val feats = features_of ctxt prover thy (Local, General) [t]
blanchet@49418
   692
          val deps = used_ths |> map nickname_of
blanchet@49418
   693
        in
blanchet@49449
   694
          mash_peek ctxt (fn {fact_G} =>
blanchet@49449
   695
              let val parents = maximal_in_graph fact_G facts in
blanchet@49449
   696
                mash_ADD ctxt overlord [(name, parents, feats, deps)]
blanchet@49449
   697
              end);
blanchet@49449
   698
          (true, "")
blanchet@49418
   699
        end)
blanchet@49398
   700
blanchet@49455
   701
val evil_theories =
blanchet@49455
   702
  ["Code_Evaluation", "DSequence", "Enum", "Lazy_Sequence", "Nitpick",
blanchet@49455
   703
   "New_DSequence", "New_Random_Sequence", "Quickcheck",
blanchet@49455
   704
   "Quickcheck_Exhaustive", "Quickcheck_Narrowing", "Random_Sequence", "SMT"]
blanchet@49455
   705
blanchet@49455
   706
fun fact_has_evil_theory (_, th) =
blanchet@49455
   707
  member (op =) evil_theories (Context.theory_name (theory_of_thm th))
blanchet@49455
   708
blanchet@49407
   709
fun sendback sub =
blanchet@49407
   710
  Markup.markup Isabelle_Markup.sendback (sledgehammerN ^ " " ^ sub)
blanchet@49407
   711
blanchet@49407
   712
val commit_timeout = seconds 30.0
blanchet@49347
   713
blanchet@49333
   714
(* The timeout is understood in a very slack fashion. *)
blanchet@49419
   715
fun mash_learn_facts ctxt (params as {debug, verbose, overlord, ...}) prover
blanchet@49419
   716
                     auto_level atp learn_timeout facts =
blanchet@49319
   717
  let
blanchet@49333
   718
    val timer = Timer.startRealTimer ()
blanchet@49407
   719
    fun next_commit_time () =
blanchet@49407
   720
      Time.+ (Timer.checkRealTimer timer, commit_timeout)
blanchet@49415
   721
    val {fact_G} = mash_get ctxt
blanchet@49415
   722
    val (old_facts, new_facts) =
blanchet@49455
   723
      facts |> filter_out fact_has_evil_theory
blanchet@49455
   724
            |> List.partition (is_fact_in_graph fact_G)
blanchet@49415
   725
            ||> sort (thm_ord o pairself snd)
blanchet@49323
   726
  in
blanchet@49419
   727
    if null new_facts andalso (not atp orelse null old_facts) then
blanchet@49419
   728
      if auto_level < 2 then
blanchet@49419
   729
        "No new " ^ (if atp then "ATP" else "Isar") ^ " proofs to learn." ^
blanchet@49419
   730
        (if auto_level = 0 andalso not atp then
blanchet@49419
   731
           "\n\nHint: Try " ^ sendback learn_atpN ^ " to learn from ATP proofs."
blanchet@49419
   732
         else
blanchet@49419
   733
           "")
blanchet@49407
   734
      else
blanchet@49407
   735
        ""
blanchet@49323
   736
    else
blanchet@49319
   737
      let
blanchet@49330
   738
        val all_names =
blanchet@49453
   739
          facts |> map snd |> map (rpair () o nickname_of) |> Symtab.make
blanchet@49454
   740
        fun deps_of status th =
blanchet@49454
   741
          if status = Non_Rec_Def orelse status = Rec_Def then
blanchet@49454
   742
            SOME []
blanchet@49454
   743
          else if atp then
blanchet@49454
   744
            atp_dependencies_of ctxt params prover auto_level facts all_names th
blanchet@49419
   745
          else
blanchet@49454
   746
            isar_dependencies_of all_names th
blanchet@49419
   747
        fun do_commit [] [] state = state
blanchet@49419
   748
          | do_commit adds reps {fact_G} =
blanchet@49407
   749
            let
blanchet@49419
   750
              val (adds, fact_G) =
blanchet@49419
   751
                ([], fact_G) |> fold (add_to_fact_graph ctxt) adds
blanchet@49419
   752
            in
blanchet@49419
   753
              mash_ADD ctxt overlord (rev adds);
blanchet@49419
   754
              mash_REPROVE ctxt overlord reps;
blanchet@49419
   755
              {fact_G = fact_G}
blanchet@49419
   756
            end
blanchet@49419
   757
        fun commit last adds reps =
blanchet@49419
   758
          (if debug andalso auto_level = 0 then
blanchet@49419
   759
             Output.urgent_message "Committing..."
blanchet@49419
   760
           else
blanchet@49419
   761
             ();
blanchet@49419
   762
           mash_map ctxt (do_commit (rev adds) reps);
blanchet@49419
   763
           if not last andalso auto_level = 0 then
blanchet@49419
   764
             let val num_proofs = length adds + length reps in
blanchet@49419
   765
               "Learned " ^ string_of_int num_proofs ^ " " ^
blanchet@49419
   766
               (if atp then "ATP" else "Isar") ^ " proof" ^
blanchet@49419
   767
               plural_s num_proofs ^ " in the last " ^
blanchet@49407
   768
               string_from_time commit_timeout ^ "."
blanchet@49407
   769
               |> Output.urgent_message
blanchet@49407
   770
             end
blanchet@49407
   771
           else
blanchet@49407
   772
             ())
blanchet@49419
   773
        fun learn_new_fact _ (accum as (_, (_, _, _, true))) = accum
blanchet@49454
   774
          | learn_new_fact ((_, stature as (_, status)), th)
blanchet@49419
   775
                           (adds, (parents, n, next_commit, _)) =
blanchet@49333
   776
            let
blanchet@49393
   777
              val name = nickname_of th
blanchet@49347
   778
              val feats =
blanchet@49400
   779
                features_of ctxt prover (theory_of_thm th) stature [prop_of th]
blanchet@49454
   780
              val deps = deps_of status th |> these
blanchet@49409
   781
              val n = n |> not (null deps) ? Integer.add 1
blanchet@49419
   782
              val adds = (name, parents, feats, deps) :: adds
blanchet@49419
   783
              val (adds, next_commit) =
blanchet@49407
   784
                if Time.> (Timer.checkRealTimer timer, next_commit) then
blanchet@49419
   785
                  (commit false adds []; ([], next_commit_time ()))
blanchet@49407
   786
                else
blanchet@49419
   787
                  (adds, next_commit)
blanchet@49419
   788
              val timed_out = Time.> (Timer.checkRealTimer timer, learn_timeout)
blanchet@49419
   789
            in (adds, ([name], n, next_commit, timed_out)) end
blanchet@49419
   790
        val n =
blanchet@49419
   791
          if null new_facts then
blanchet@49419
   792
            0
blanchet@49419
   793
          else
blanchet@49419
   794
            let
blanchet@49419
   795
              val last_th = new_facts |> List.last |> snd
blanchet@49419
   796
              (* crude approximation *)
blanchet@49419
   797
              val ancestors =
blanchet@49419
   798
                old_facts
blanchet@49419
   799
                |> filter (fn (_, th) => thm_ord (th, last_th) <> GREATER)
blanchet@49422
   800
              val parents = maximal_in_graph fact_G ancestors
blanchet@49419
   801
              val (adds, (_, n, _, _)) =
blanchet@49419
   802
                ([], (parents, 0, next_commit_time (), false))
blanchet@49419
   803
                |> fold learn_new_fact new_facts
blanchet@49419
   804
            in commit true adds []; n end
blanchet@49419
   805
        fun relearn_old_fact _ (accum as (_, (_, _, true))) = accum
blanchet@49454
   806
          | relearn_old_fact ((_, (_, status)), th)
blanchet@49454
   807
                             (reps, (n, next_commit, _)) =
blanchet@49419
   808
            let
blanchet@49419
   809
              val name = nickname_of th
blanchet@49419
   810
              val (n, reps) =
blanchet@49454
   811
                case deps_of status th of
blanchet@49419
   812
                  SOME deps => (n + 1, (name, deps) :: reps)
blanchet@49419
   813
                | NONE => (n, reps)
blanchet@49419
   814
              val (reps, next_commit) =
blanchet@49419
   815
                if Time.> (Timer.checkRealTimer timer, next_commit) then
blanchet@49419
   816
                  (commit false [] reps; ([], next_commit_time ()))
blanchet@49419
   817
                else
blanchet@49419
   818
                  (reps, next_commit)
blanchet@49419
   819
              val timed_out = Time.> (Timer.checkRealTimer timer, learn_timeout)
blanchet@49419
   820
            in (reps, (n, next_commit, timed_out)) end
blanchet@49419
   821
        val n =
blanchet@49448
   822
          if not atp orelse null old_facts then
blanchet@49419
   823
            n
blanchet@49419
   824
          else
blanchet@49419
   825
            let
blanchet@49421
   826
              fun priority_of (_, th) =
blanchet@49419
   827
                random_range 0 (1000 * max_dependencies)
blanchet@49419
   828
                - 500 * (th |> isar_dependencies_of all_names
blanchet@49419
   829
                            |> Option.map length
blanchet@49419
   830
                            |> the_default max_dependencies)
blanchet@49419
   831
              val old_facts =
blanchet@49421
   832
                old_facts |> map (`priority_of)
blanchet@49419
   833
                          |> sort (int_ord o pairself fst)
blanchet@49419
   834
                          |> map snd
blanchet@49419
   835
              val (reps, (n, _, _)) =
blanchet@49419
   836
                ([], (n, next_commit_time (), false))
blanchet@49419
   837
                |> fold relearn_old_fact old_facts
blanchet@49419
   838
            in commit true [] reps; n end
blanchet@49333
   839
      in
blanchet@49419
   840
        if verbose orelse auto_level < 2 then
blanchet@49419
   841
          "Learned " ^ string_of_int n ^ " nontrivial " ^
blanchet@49419
   842
          (if atp then "ATP" else "Isar") ^ " proof" ^ plural_s n ^
blanchet@49334
   843
          (if verbose then
blanchet@49334
   844
             " in " ^ string_from_time (Timer.checkRealTimer timer)
blanchet@49334
   845
           else
blanchet@49334
   846
             "") ^ "."
blanchet@49334
   847
        else
blanchet@49334
   848
          ""
blanchet@49333
   849
      end
blanchet@49323
   850
  end
blanchet@49319
   851
blanchet@49419
   852
fun mash_learn ctxt (params as {provers, timeout, ...}) fact_override chained
blanchet@49419
   853
               atp =
blanchet@49331
   854
  let
blanchet@49411
   855
    val css = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
blanchet@49410
   856
    val ctxt = ctxt |> Config.put instantiate_inducts false
blanchet@49410
   857
    val facts =
blanchet@49411
   858
      nearly_all_facts ctxt false fact_override Symtab.empty css chained []
blanchet@49411
   859
                       @{prop True}
blanchet@49419
   860
    val num_facts = length facts
blanchet@49419
   861
    val prover = hd provers
blanchet@49419
   862
    fun learn auto_level atp =
blanchet@49419
   863
      mash_learn_facts ctxt params prover auto_level atp infinite_timeout facts
blanchet@49419
   864
      |> Output.urgent_message
blanchet@49331
   865
  in
blanchet@49419
   866
    (if atp then
blanchet@49419
   867
       ("MaShing through " ^ string_of_int num_facts ^ " fact" ^
blanchet@49419
   868
        plural_s num_facts ^ " for ATP proofs (" ^ quote prover ^ " timeout: " ^
blanchet@49419
   869
        string_from_time timeout ^ ").\n\nCollecting Isar proofs first..."
blanchet@49419
   870
        |> Output.urgent_message;
blanchet@49419
   871
        learn 1 false;
blanchet@49419
   872
        "Now collecting ATP proofs. This may take several hours. You can \
blanchet@49419
   873
        \safely stop the learning process at any point."
blanchet@49419
   874
        |> Output.urgent_message;
blanchet@49419
   875
        learn 0 true)
blanchet@49419
   876
     else
blanchet@49419
   877
       ("MaShing through " ^ string_of_int num_facts ^ " fact" ^
blanchet@49419
   878
        plural_s num_facts ^ " for Isar proofs..."
blanchet@49419
   879
        |> Output.urgent_message;
blanchet@49419
   880
        learn 0 false))
blanchet@49331
   881
  end
blanchet@49264
   882
blanchet@49333
   883
(* The threshold should be large enough so that MaSh doesn't kick in for Auto
blanchet@49333
   884
   Sledgehammer and Try. *)
blanchet@49333
   885
val min_secs_for_learning = 15
blanchet@49333
   886
blanchet@49336
   887
fun relevant_facts ctxt (params as {learn, fact_filter, timeout, ...}) prover
blanchet@49336
   888
        max_facts ({add, only, ...} : fact_override) hyp_ts concl_t facts =
blanchet@49329
   889
  if not (subset (op =) (the_list fact_filter, fact_filters)) then
blanchet@49329
   890
    error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".")
blanchet@49329
   891
  else if only then
blanchet@49304
   892
    facts
blanchet@49336
   893
  else if max_facts <= 0 orelse null facts then
blanchet@49303
   894
    []
blanchet@49303
   895
  else
blanchet@49303
   896
    let
blanchet@49342
   897
      fun maybe_learn () =
blanchet@49399
   898
        if learn andalso not (Async_Manager.has_running_threads MaShN) andalso
blanchet@49399
   899
           Time.toSeconds timeout >= min_secs_for_learning then
blanchet@49399
   900
          let val timeout = time_mult learn_timeout_slack timeout in
blanchet@49399
   901
            launch_thread timeout
blanchet@49419
   902
                (fn () => (true, mash_learn_facts ctxt params prover 2 false
blanchet@49407
   903
                                                  timeout facts))
blanchet@49334
   904
          end
blanchet@49333
   905
        else
blanchet@49333
   906
          ()
blanchet@49329
   907
      val fact_filter =
blanchet@49329
   908
        case fact_filter of
blanchet@49394
   909
          SOME ff => (() |> ff <> mepoN ? maybe_learn; ff)
blanchet@49333
   910
        | NONE =>
blanchet@49422
   911
          if is_smt_prover ctxt prover then
blanchet@49422
   912
            mepoN
blanchet@49422
   913
          else if mash_could_suggest_facts () then
blanchet@49422
   914
            (maybe_learn ();
blanchet@49422
   915
             if mash_can_suggest_facts ctxt then meshN else mepoN)
blanchet@49422
   916
          else
blanchet@49422
   917
            mepoN
blanchet@49303
   918
      val add_ths = Attrib.eval_thms ctxt add
blanchet@49307
   919
      fun prepend_facts ths accepts =
blanchet@49303
   920
        ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
blanchet@49307
   921
         (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
blanchet@49308
   922
        |> take max_facts
blanchet@49421
   923
      fun mepo () =
blanchet@49421
   924
        facts |> mepo_suggested_facts ctxt params prover max_facts NONE hyp_ts
blanchet@49421
   925
                                      concl_t
blanchet@49421
   926
              |> weight_mepo_facts
blanchet@49329
   927
      fun mash () =
blanchet@49421
   928
        mash_suggested_facts ctxt params prover max_facts hyp_ts concl_t facts
blanchet@49329
   929
      val mess =
blanchet@49421
   930
        [] |> (if fact_filter <> mashN then cons (mepo (), []) else I)
blanchet@49394
   931
           |> (if fact_filter <> mepoN then cons (mash ()) else I)
blanchet@49303
   932
    in
blanchet@49328
   933
      mesh_facts max_facts mess
blanchet@49303
   934
      |> not (null add_ths) ? prepend_facts add_ths
blanchet@49303
   935
    end
blanchet@49303
   936
blanchet@49334
   937
fun kill_learners () = Async_Manager.kill_threads MaShN "learner"
blanchet@49334
   938
fun running_learners () = Async_Manager.running_threads MaShN "learner"
blanchet@49334
   939
blanchet@49263
   940
end;