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