src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 49333 325c8fd0d762
parent 49331 252f45c04042
child 49334 340187063d84
permissions -rw-r--r--
more consolidation of MaSh code
blanchet@49263
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
blanchet@49263
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@49263
     3
blanchet@49263
     4
Sledgehammer's machine-learning-based relevance filter (MaSh).
blanchet@49263
     5
*)
blanchet@49263
     6
blanchet@49263
     7
signature SLEDGEHAMMER_FILTER_MASH =
blanchet@49263
     8
sig
blanchet@49266
     9
  type status = ATP_Problem_Generate.status
blanchet@49266
    10
  type stature = ATP_Problem_Generate.stature
blanchet@49311
    11
  type fact = Sledgehammer_Fact.fact
blanchet@49311
    12
  type fact_override = Sledgehammer_Fact.fact_override
blanchet@49266
    13
  type params = Sledgehammer_Provers.params
blanchet@49303
    14
  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
blanchet@49266
    15
  type prover_result = Sledgehammer_Provers.prover_result
blanchet@49266
    16
blanchet@49323
    17
  val trace : bool Config.T
blanchet@49329
    18
  val meshN : string
blanchet@49329
    19
  val iterN : string
blanchet@49329
    20
  val mashN : 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@49327
    27
  val suggested_facts : string list -> fact list -> fact list
blanchet@49328
    28
  val mesh_facts : int -> (fact list * int option) list -> fact list
blanchet@49333
    29
  val is_likely_tautology : Proof.context -> string -> thm -> bool
blanchet@49330
    30
  val is_too_meta : thm -> bool
blanchet@49266
    31
  val theory_ord : theory * theory -> order
blanchet@49266
    32
  val thm_ord : thm * thm -> order
blanchet@49333
    33
  val features_of :
blanchet@49333
    34
    Proof.context -> string -> theory -> status -> term list -> string list
blanchet@49331
    35
  val isabelle_dependencies_of : unit Symtab.table -> thm -> string list
blanchet@49266
    36
  val goal_of_thm : theory -> thm -> thm
blanchet@49333
    37
  val run_prover :
blanchet@49333
    38
    Proof.context -> params -> string -> fact list -> thm -> prover_result
blanchet@49323
    39
  val mash_RESET : Proof.context -> unit
blanchet@49331
    40
  val mash_INIT :
blanchet@49331
    41
    Proof.context -> bool
blanchet@49331
    42
    -> (string * string list * string list * string list) list -> unit
blanchet@49323
    43
  val mash_ADD :
blanchet@49331
    44
    Proof.context -> bool
blanchet@49331
    45
    -> (string * string list * string list * string list) list -> unit
blanchet@49331
    46
  val mash_QUERY :
blanchet@49333
    47
    Proof.context -> bool -> int -> string list * string list -> string list
blanchet@49323
    48
  val mash_reset : Proof.context -> unit
blanchet@49333
    49
  val mash_could_suggest_facts : unit -> bool
blanchet@49333
    50
  val mash_can_suggest_facts : unit -> bool
blanchet@49313
    51
  val mash_suggest_facts :
blanchet@49333
    52
    Proof.context -> params -> string -> int -> term list -> term -> fact list
blanchet@49326
    53
    -> fact list
blanchet@49333
    54
  val mash_learn_thy : Proof.context -> params -> theory -> Time.time -> unit
blanchet@49331
    55
  val mash_learn_proof :
blanchet@49331
    56
    Proof.context -> params -> term -> thm list -> fact list -> unit
blanchet@49303
    57
  val relevant_facts :
blanchet@49307
    58
    Proof.context -> params -> string -> int -> fact_override -> term list
blanchet@49311
    59
    -> term -> fact list -> fact list
blanchet@49263
    60
end;
blanchet@49263
    61
blanchet@49263
    62
structure Sledgehammer_Filter_MaSh : SLEDGEHAMMER_FILTER_MASH =
blanchet@49263
    63
struct
blanchet@49264
    64
blanchet@49266
    65
open ATP_Util
blanchet@49266
    66
open ATP_Problem_Generate
blanchet@49266
    67
open Sledgehammer_Util
blanchet@49266
    68
open Sledgehammer_Fact
blanchet@49266
    69
open Sledgehammer_Filter_Iter
blanchet@49266
    70
open Sledgehammer_Provers
blanchet@49333
    71
open Sledgehammer_Minimize
blanchet@49266
    72
blanchet@49323
    73
val trace =
blanchet@49323
    74
  Attrib.setup_config_bool @{binding sledgehammer_filter_mash_trace} (K false)
blanchet@49323
    75
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
blanchet@49323
    76
blanchet@49329
    77
val meshN = "mesh"
blanchet@49329
    78
val iterN = "iter"
blanchet@49329
    79
val mashN = "mash"
blanchet@49329
    80
blanchet@49329
    81
val fact_filters = [meshN, iterN, mashN]
blanchet@49329
    82
blanchet@49329
    83
fun mash_home () = getenv "MASH_HOME"
blanchet@49329
    84
fun mash_state_dir () =
blanchet@49324
    85
  getenv "ISABELLE_HOME_USER" ^ "/mash"
blanchet@49331
    86
  |> tap (Isabelle_System.mkdir o Path.explode)
blanchet@49329
    87
fun mash_state_path () = mash_state_dir () ^ "/state" |> Path.explode
blanchet@49266
    88
blanchet@49266
    89
(*** Isabelle helpers ***)
blanchet@49266
    90
blanchet@49323
    91
fun meta_char c =
blanchet@49266
    92
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
blanchet@49266
    93
     c = #")" orelse c = #"," then
blanchet@49266
    94
    String.str c
blanchet@49266
    95
  else
blanchet@49266
    96
    (* fixed width, in case more digits follow *)
blanchet@49266
    97
    "\\" ^ stringN_of_int 3 (Char.ord c)
blanchet@49266
    98
blanchet@49323
    99
fun unmeta_chars accum [] = String.implode (rev accum)
blanchet@49323
   100
  | unmeta_chars accum (#"\\" :: d1 :: d2 :: d3 :: cs) =
blanchet@49323
   101
    (case Int.fromString (String.implode [d1, d2, d3]) of
blanchet@49323
   102
       SOME n => unmeta_chars (Char.chr n :: accum) cs
blanchet@49323
   103
     | NONE => "" (* error *))
blanchet@49323
   104
  | unmeta_chars _ (#"\\" :: _) = "" (* error *)
blanchet@49323
   105
  | unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs
blanchet@49323
   106
blanchet@49323
   107
val escape_meta = String.translate meta_char
blanchet@49318
   108
val escape_metas = map escape_meta #> space_implode " "
blanchet@49330
   109
val unescape_meta = String.explode #> unmeta_chars []
blanchet@49330
   110
val unescape_metas =
blanchet@49330
   111
  space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta
blanchet@49266
   112
blanchet@49326
   113
fun extract_query line =
blanchet@49326
   114
  case space_explode ":" line of
blanchet@49330
   115
    [goal_name, suggs] => (unescape_meta goal_name, unescape_metas suggs)
blanchet@49327
   116
  | _ => ("", [])
blanchet@49326
   117
blanchet@49326
   118
fun find_suggested facts sugg =
blanchet@49326
   119
  find_first (fn (_, th) => Thm.get_name_hint th = sugg) facts
blanchet@49326
   120
fun suggested_facts suggs facts = map_filter (find_suggested facts) suggs
blanchet@49326
   121
blanchet@49328
   122
fun sum_avg n xs =
blanchet@49328
   123
  fold (Integer.add o Integer.mult n) xs 0 div (length xs)
blanchet@49328
   124
blanchet@49329
   125
fun mesh_facts max_facts [(facts, _)] = facts |> take max_facts
blanchet@49329
   126
  | mesh_facts max_facts mess =
blanchet@49329
   127
    let
blanchet@49329
   128
      val n = length mess
blanchet@49329
   129
      val fact_eq = Thm.eq_thm o pairself snd
blanchet@49329
   130
      fun score_in fact (facts, def) =
blanchet@49329
   131
        case find_index (curry fact_eq fact) facts of
blanchet@49329
   132
          ~1 => def
blanchet@49329
   133
        | j => SOME j
blanchet@49329
   134
      fun score_of fact = mess |> map_filter (score_in fact) |> sum_avg n
blanchet@49329
   135
      val facts = fold (union fact_eq o take max_facts o fst) mess []
blanchet@49329
   136
    in
blanchet@49329
   137
      facts |> map (`score_of) |> sort (int_ord o pairself fst) |> map snd
blanchet@49329
   138
            |> take max_facts
blanchet@49329
   139
    end
blanchet@49327
   140
blanchet@49318
   141
val thy_feature_prefix = "y_"
blanchet@49266
   142
blanchet@49318
   143
val thy_feature_name_of = prefix thy_feature_prefix
blanchet@49318
   144
val const_name_of = prefix const_prefix
blanchet@49318
   145
val type_name_of = prefix type_const_prefix
blanchet@49318
   146
val class_name_of = prefix class_prefix
blanchet@49266
   147
blanchet@49266
   148
local
blanchet@49266
   149
blanchet@49266
   150
fun has_bool @{typ bool} = true
blanchet@49266
   151
  | has_bool (Type (_, Ts)) = exists has_bool Ts
blanchet@49266
   152
  | has_bool _ = false
blanchet@49266
   153
blanchet@49266
   154
fun has_fun (Type (@{type_name fun}, _)) = true
blanchet@49266
   155
  | has_fun (Type (_, Ts)) = exists has_fun Ts
blanchet@49266
   156
  | has_fun _ = false
blanchet@49266
   157
blanchet@49266
   158
val is_conn = member (op =)
blanchet@49266
   159
  [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
blanchet@49266
   160
   @{const_name HOL.implies}, @{const_name Not},
blanchet@49266
   161
   @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex},
blanchet@49266
   162
   @{const_name HOL.eq}]
blanchet@49266
   163
blanchet@49266
   164
val has_bool_arg_const =
blanchet@49266
   165
  exists_Const (fn (c, T) =>
blanchet@49266
   166
                   not (is_conn c) andalso exists has_bool (binder_types T))
blanchet@49266
   167
blanchet@49331
   168
fun higher_inst_const thy (s, T) =
blanchet@49266
   169
  case binder_types T of
blanchet@49266
   170
    [] => false
blanchet@49331
   171
  | Ts => length (binder_types (Sign.the_const_type thy s)) <> length Ts
blanchet@49331
   172
  handle TYPE _ => false
blanchet@49266
   173
blanchet@49266
   174
val binders = [@{const_name All}, @{const_name Ex}]
blanchet@49266
   175
blanchet@49266
   176
in
blanchet@49266
   177
blanchet@49266
   178
fun is_fo_term thy t =
blanchet@49266
   179
  let
blanchet@49266
   180
    val t =
blanchet@49266
   181
      t |> Envir.beta_eta_contract
blanchet@49266
   182
        |> transform_elim_prop
blanchet@49266
   183
        |> Object_Logic.atomize_term thy
blanchet@49266
   184
  in
blanchet@49266
   185
    Term.is_first_order binders t andalso
blanchet@49266
   186
    not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T
blanchet@49266
   187
                          | _ => false) t orelse
blanchet@49266
   188
         has_bool_arg_const t orelse exists_Const (higher_inst_const thy) t)
blanchet@49266
   189
  end
blanchet@49266
   190
blanchet@49266
   191
end
blanchet@49266
   192
blanchet@49333
   193
fun interesting_terms_types_and_classes ctxt prover term_max_depth
blanchet@49333
   194
                                        type_max_depth ts =
blanchet@49266
   195
  let
blanchet@49266
   196
    val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49333
   197
    fun is_bad_const (x as (s, _)) args =
blanchet@49333
   198
      member (op =) atp_logical_consts s orelse
blanchet@49333
   199
      fst (is_built_in_const_for_prover ctxt prover x args)
blanchet@49319
   200
    fun add_classes @{sort type} = I
blanchet@49319
   201
      | add_classes S = union (op =) (map class_name_of S)
blanchet@49266
   202
    fun do_add_type (Type (s, Ts)) =
blanchet@49266
   203
        (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
blanchet@49266
   204
        #> fold do_add_type Ts
blanchet@49319
   205
      | do_add_type (TFree (_, S)) = add_classes S
blanchet@49319
   206
      | do_add_type (TVar (_, S)) = add_classes S
blanchet@49266
   207
    fun add_type T = type_max_depth >= 0 ? do_add_type T
blanchet@49266
   208
    fun mk_app s args =
blanchet@49266
   209
      if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
blanchet@49266
   210
      else s
blanchet@49266
   211
    fun patternify ~1 _ = ""
blanchet@49266
   212
      | patternify depth t =
blanchet@49266
   213
        case strip_comb t of
blanchet@49266
   214
          (Const (s, _), args) =>
blanchet@49266
   215
          mk_app (const_name_of s) (map (patternify (depth - 1)) args)
blanchet@49266
   216
        | _ => ""
blanchet@49266
   217
    fun add_term_patterns ~1 _ = I
blanchet@49266
   218
      | add_term_patterns depth t =
blanchet@49266
   219
        insert (op =) (patternify depth t)
blanchet@49266
   220
        #> add_term_patterns (depth - 1) t
blanchet@49266
   221
    val add_term = add_term_patterns term_max_depth
blanchet@49266
   222
    fun add_patterns t =
blanchet@49266
   223
      let val (head, args) = strip_comb t in
blanchet@49266
   224
        (case head of
blanchet@49333
   225
           Const (x as (_, T)) =>
blanchet@49333
   226
           not (is_bad_const x args) ? (add_term t #> add_type T)
blanchet@49266
   227
         | Free (_, T) => add_type T
blanchet@49266
   228
         | Var (_, T) => add_type T
blanchet@49266
   229
         | Abs (_, T, body) => add_type T #> add_patterns body
blanchet@49266
   230
         | _ => I)
blanchet@49266
   231
        #> fold add_patterns args
blanchet@49266
   232
      end
blanchet@49317
   233
  in [] |> fold add_patterns ts |> sort string_ord end
blanchet@49266
   234
blanchet@49333
   235
fun is_likely_tautology ctxt prover th =
blanchet@49333
   236
  null (interesting_terms_types_and_classes ctxt prover 0 ~1 [prop_of th])
blanchet@49333
   237
  andalso not (Thm.eq_thm_prop (@{thm ext}, th))
blanchet@49266
   238
blanchet@49330
   239
(* ### FIXME: optimize *)
blanchet@49330
   240
fun is_too_meta th =
blanchet@49330
   241
  fastype_of (Object_Logic.atomize_term (theory_of_thm th) (prop_of th))
blanchet@49330
   242
  <> @{typ bool}
blanchet@49266
   243
blanchet@49266
   244
fun theory_ord p =
blanchet@49330
   245
  if Theory.eq_thy p then
blanchet@49330
   246
    EQUAL
blanchet@49330
   247
  else if Theory.subthy p then
blanchet@49330
   248
    LESS
blanchet@49330
   249
  else if Theory.subthy (swap p) then
blanchet@49330
   250
    GREATER
blanchet@49330
   251
  else case int_ord (pairself (length o Theory.ancestors_of) p) of
blanchet@49330
   252
    EQUAL => string_ord (pairself Context.theory_name p)
blanchet@49330
   253
  | order => order
blanchet@49266
   254
blanchet@49266
   255
val thm_ord = theory_ord o pairself theory_of_thm
blanchet@49266
   256
blanchet@49266
   257
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
blanchet@49266
   258
blanchet@49312
   259
val term_max_depth = 1
blanchet@49312
   260
val type_max_depth = 1
blanchet@49266
   261
blanchet@49266
   262
(* TODO: Generate type classes for types? *)
blanchet@49333
   263
fun features_of ctxt prover thy status ts =
blanchet@49318
   264
  thy_feature_name_of (Context.theory_name thy) ::
blanchet@49333
   265
  interesting_terms_types_and_classes ctxt prover term_max_depth type_max_depth
blanchet@49333
   266
                                      ts
blanchet@49317
   267
  |> exists (not o is_lambda_free) ts ? cons "lambdas"
blanchet@49317
   268
  |> exists (exists_Const is_exists) ts ? cons "skolems"
blanchet@49317
   269
  |> exists (not o is_fo_term thy) ts ? cons "ho"
blanchet@49317
   270
  |> (case status of
blanchet@49317
   271
        General => I
blanchet@49317
   272
      | Induction => cons "induction"
blanchet@49317
   273
      | Intro => cons "intro"
blanchet@49317
   274
      | Inductive => cons "inductive"
blanchet@49317
   275
      | Elim => cons "elim"
blanchet@49317
   276
      | Simp => cons "simp"
blanchet@49317
   277
      | Def => cons "def")
blanchet@49266
   278
blanchet@49266
   279
fun isabelle_dependencies_of all_facts =
blanchet@49318
   280
  thms_in_proof (SOME all_facts) #> sort string_ord
blanchet@49266
   281
blanchet@49266
   282
val freezeT = Type.legacy_freeze_type
blanchet@49266
   283
blanchet@49266
   284
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49266
   285
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49266
   286
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49266
   287
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49266
   288
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49266
   289
  | freeze t = t
blanchet@49266
   290
blanchet@49266
   291
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
blanchet@49266
   292
blanchet@49333
   293
fun run_prover ctxt params prover facts goal =
blanchet@49266
   294
  let
blanchet@49266
   295
    val problem =
blanchet@49266
   296
      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
blanchet@49304
   297
       facts = facts |> map (apfst (apfst (fn name => name ())))
blanchet@49333
   298
                     |> map Untranslated_Fact}
blanchet@49333
   299
    val prover = get_minimizing_prover ctxt Normal prover
blanchet@49266
   300
  in prover params (K (K (K ""))) problem end
blanchet@49266
   301
blanchet@49266
   302
blanchet@49317
   303
(*** Low-level communication with MaSh ***)
blanchet@49317
   304
blanchet@49333
   305
fun write_file write file =
blanchet@49333
   306
  let val path = Path.explode file in
blanchet@49333
   307
    File.write path ""; write (File.append path)
blanchet@49333
   308
  end
blanchet@49331
   309
blanchet@49331
   310
fun mash_info overlord =
blanchet@49331
   311
  if overlord then (getenv "ISABELLE_HOME_USER", "")
blanchet@49331
   312
  else (getenv "ISABELLE_TMP", serial_string ())
blanchet@49331
   313
blanchet@49331
   314
fun run_mash ctxt (temp_dir, serial) core =
blanchet@49326
   315
  let
blanchet@49331
   316
    val log_file = temp_dir ^ "/mash_log" ^ serial
blanchet@49331
   317
    val err_file = temp_dir ^ "/mash_err" ^ serial
blanchet@49326
   318
    val command =
blanchet@49331
   319
      mash_home () ^ "/mash.py --quiet --outputDir " ^ mash_state_dir () ^
blanchet@49331
   320
      " --log " ^ log_file ^ " " ^ core ^ " 2>&1 > " ^ err_file
blanchet@49331
   321
  in
blanchet@49331
   322
    trace_msg ctxt (fn () => "Running " ^ command);
blanchet@49333
   323
    write_file (K ()) log_file;
blanchet@49333
   324
    write_file (K ()) err_file;
blanchet@49331
   325
    Isabelle_System.bash command; ()
blanchet@49331
   326
  end
blanchet@49326
   327
blanchet@49331
   328
fun run_mash_init ctxt overlord write_access write_feats write_deps =
blanchet@49331
   329
  let
blanchet@49331
   330
    val info as (temp_dir, serial) = mash_info overlord
blanchet@49331
   331
    val in_dir = temp_dir ^ "/mash_init" ^ serial
blanchet@49331
   332
                 |> tap (Isabelle_System.mkdir o Path.explode)
blanchet@49331
   333
  in
blanchet@49331
   334
    write_file write_access (in_dir ^ "/mash_accessibility");
blanchet@49331
   335
    write_file write_feats (in_dir ^ "/mash_features");
blanchet@49331
   336
    write_file write_deps (in_dir ^ "/mash_dependencies");
blanchet@49331
   337
    run_mash ctxt info ("--init --inputDir " ^ in_dir)
blanchet@49331
   338
  end
blanchet@49331
   339
blanchet@49333
   340
fun run_mash_commands ctxt overlord save max_suggs write_cmds read_suggs =
blanchet@49331
   341
  let
blanchet@49331
   342
    val info as (temp_dir, serial) = mash_info overlord
blanchet@49333
   343
    val sugg_file = temp_dir ^ "/mash_suggs" ^ serial
blanchet@49331
   344
    val cmd_file = temp_dir ^ "/mash_commands" ^ serial
blanchet@49331
   345
  in
blanchet@49333
   346
    write_file (K ()) sugg_file;
blanchet@49331
   347
    write_file write_cmds cmd_file;
blanchet@49331
   348
    run_mash ctxt info
blanchet@49333
   349
             ("--inputFile " ^ cmd_file ^ " --predictions " ^ sugg_file ^
blanchet@49333
   350
              " --numberOfPredictions " ^ string_of_int max_suggs ^
blanchet@49331
   351
              (if save then " --saveModel" else ""));
blanchet@49333
   352
    read_suggs (fn () => File.read_lines (Path.explode sugg_file))
blanchet@49331
   353
  end
blanchet@49331
   354
blanchet@49331
   355
fun str_of_update (name, parents, feats, deps) =
blanchet@49331
   356
  "! " ^ escape_meta name ^ ": " ^ escape_metas parents ^ "; " ^
blanchet@49326
   357
  escape_metas feats ^ "; " ^ escape_metas deps ^ "\n"
blanchet@49326
   358
blanchet@49331
   359
fun str_of_query (parents, feats) =
blanchet@49331
   360
  "? " ^ escape_metas parents ^ "; " ^ escape_metas feats
blanchet@49331
   361
blanchet@49331
   362
fun init_str_of_update get (upd as (name, _, _, _)) =
blanchet@49331
   363
  escape_meta name ^ ": " ^ escape_metas (get upd) ^ "\n"
blanchet@49326
   364
blanchet@49323
   365
fun mash_RESET ctxt =
blanchet@49329
   366
  let val path = mash_state_dir () |> Path.explode in
blanchet@49324
   367
    trace_msg ctxt (K "MaSh RESET");
blanchet@49324
   368
    File.fold_dir (fn file => fn () =>
blanchet@49324
   369
                      File.rm (Path.append path (Path.basic file)))
blanchet@49324
   370
                  path ()
blanchet@49324
   371
  end
blanchet@49317
   372
blanchet@49331
   373
fun mash_INIT ctxt _ [] = mash_RESET ctxt
blanchet@49331
   374
  | mash_INIT ctxt overlord upds =
blanchet@49331
   375
    (trace_msg ctxt (fn () => "MaSh INIT " ^
blanchet@49331
   376
         elide_string 1000 (space_implode " " (map #1 upds)));
blanchet@49331
   377
     run_mash_init ctxt overlord
blanchet@49331
   378
         (fn append => List.app (append o init_str_of_update #2) upds)
blanchet@49331
   379
         (fn append => List.app (append o init_str_of_update #3) upds)
blanchet@49331
   380
         (fn append => List.app (append o init_str_of_update #4) upds))
blanchet@49317
   381
blanchet@49331
   382
fun mash_ADD _ _ [] = ()
blanchet@49331
   383
  | mash_ADD ctxt overlord upds =
blanchet@49331
   384
    (trace_msg ctxt (fn () => "MaSh ADD " ^
blanchet@49331
   385
         elide_string 1000 (space_implode " " (map #1 upds)));
blanchet@49333
   386
     run_mash_commands ctxt overlord true 0
blanchet@49331
   387
         (fn append => List.app (append o str_of_update) upds) (K ()))
blanchet@49317
   388
blanchet@49333
   389
fun mash_QUERY ctxt overlord max_suggs (query as (_, feats)) =
blanchet@49329
   390
  (trace_msg ctxt (fn () => "MaSh QUERY " ^ space_implode " " feats);
blanchet@49333
   391
   run_mash_commands ctxt overlord false max_suggs
blanchet@49331
   392
       (fn append => append (str_of_query query))
blanchet@49333
   393
       (fn suggs => snd (extract_query (List.last (suggs ()))))
blanchet@49326
   394
   handle List.Empty => [])
blanchet@49317
   395
blanchet@49317
   396
blanchet@49266
   397
(*** High-level communication with MaSh ***)
blanchet@49266
   398
blanchet@49316
   399
type mash_state =
blanchet@49331
   400
  {thys : bool Symtab.table,
blanchet@49331
   401
   fact_graph : unit Graph.T}
blanchet@49264
   402
blanchet@49331
   403
val empty_state = {thys = Symtab.empty, fact_graph = Graph.empty}
blanchet@49316
   404
blanchet@49316
   405
local
blanchet@49316
   406
blanchet@49316
   407
fun mash_load (state as (true, _)) = state
blanchet@49316
   408
  | mash_load _ =
blanchet@49324
   409
    let val path = mash_state_path () in
blanchet@49317
   410
      (true,
blanchet@49317
   411
       case try File.read_lines path of
blanchet@49331
   412
         SOME (comp_thys :: incomp_thys :: fact_lines) =>
blanchet@49317
   413
         let
blanchet@49331
   414
           fun add_thy comp thy = Symtab.update (thy, comp)
blanchet@49331
   415
           fun add_fact_line line =
blanchet@49331
   416
             case extract_query line of
blanchet@49331
   417
               ("", _) => I (* shouldn't happen *)
blanchet@49331
   418
             | (name, parents) =>
blanchet@49331
   419
               Graph.default_node (name, ())
blanchet@49331
   420
               #> fold (fn par => Graph.add_edge (par, name)) parents
blanchet@49331
   421
           val thys =
blanchet@49331
   422
             Symtab.empty |> fold (add_thy true) (unescape_metas comp_thys)
blanchet@49331
   423
                          |> fold (add_thy false) (unescape_metas incomp_thys)
blanchet@49331
   424
           val fact_graph = Graph.empty |> fold add_fact_line fact_lines
blanchet@49331
   425
         in {thys = thys, fact_graph = fact_graph} end
blanchet@49319
   426
       | _ => empty_state)
blanchet@49317
   427
    end
blanchet@49316
   428
blanchet@49331
   429
fun mash_save ({thys, fact_graph, ...} : mash_state) =
blanchet@49316
   430
  let
blanchet@49324
   431
    val path = mash_state_path ()
blanchet@49331
   432
    val thys = Symtab.dest thys
blanchet@49333
   433
    val line_for_thys = escape_metas o AList.find (op =) thys
blanchet@49333
   434
    fun fact_line_for name parents =
blanchet@49333
   435
      escape_meta name ^ ": " ^ escape_metas parents
blanchet@49331
   436
    val append_fact = File.append path o suffix "\n" oo fact_line_for
blanchet@49316
   437
  in
blanchet@49331
   438
    File.write path (line_for_thys true ^ "\n" ^ line_for_thys false ^ "\n");
blanchet@49331
   439
    Graph.fold (fn (name, ((), (parents, _))) => fn () =>
blanchet@49331
   440
                   append_fact name (Graph.Keys.dest parents))
blanchet@49331
   441
        fact_graph ()
blanchet@49316
   442
  end
blanchet@49316
   443
blanchet@49317
   444
val global_state =
blanchet@49319
   445
  Synchronized.var "Sledgehammer_Filter_MaSh.global_state" (false, empty_state)
blanchet@49316
   446
blanchet@49316
   447
in
blanchet@49316
   448
blanchet@49321
   449
fun mash_map f =
blanchet@49317
   450
  Synchronized.change global_state (mash_load ##> (f #> tap mash_save))
blanchet@49316
   451
blanchet@49319
   452
fun mash_get () = Synchronized.change_result global_state (mash_load #> `snd)
blanchet@49317
   453
blanchet@49323
   454
fun mash_reset ctxt =
blanchet@49317
   455
  Synchronized.change global_state (fn _ =>
blanchet@49324
   456
      (mash_RESET ctxt; File.write (mash_state_path ()) "";
blanchet@49323
   457
       (true, empty_state)))
blanchet@49316
   458
blanchet@49316
   459
end
blanchet@49316
   460
blanchet@49333
   461
fun mash_could_suggest_facts () = mash_home () <> ""
blanchet@49333
   462
fun mash_can_suggest_facts () = not (Graph.is_empty (#fact_graph (mash_get ())))
blanchet@49264
   463
blanchet@49333
   464
fun parents_wrt_facts facts fact_graph =
blanchet@49331
   465
  let
blanchet@49333
   466
    val graph_facts = Symtab.make (map (rpair ()) (Graph.keys fact_graph))
blanchet@49331
   467
    val facts =
blanchet@49333
   468
      [] |> fold (cons o Thm.get_name_hint o snd) facts
blanchet@49333
   469
         |> filter (Symtab.defined graph_facts)
blanchet@49333
   470
         |> Graph.all_preds fact_graph
blanchet@49333
   471
    val facts = Symtab.empty |> fold (fn name => Symtab.update (name, ())) facts
blanchet@49331
   472
  in fact_graph |> Graph.restrict (Symtab.defined facts) |> Graph.maximals end
blanchet@49331
   473
blanchet@49333
   474
(* Generate more suggestions than requested, because some might be thrown out
blanchet@49333
   475
   later for various reasons and "meshing" gives better results with some
blanchet@49333
   476
   slack. *)
blanchet@49333
   477
fun max_suggs_of max_facts = max_facts + Int.min (200, max_facts)
blanchet@49333
   478
blanchet@49333
   479
fun mash_suggest_facts ctxt ({overlord, ...} : params) prover max_facts hyp_ts
blanchet@49333
   480
                       concl_t facts =
blanchet@49316
   481
  let
blanchet@49317
   482
    val thy = Proof_Context.theory_of ctxt
blanchet@49331
   483
    val fact_graph = #fact_graph (mash_get ())
blanchet@49333
   484
val _ = warning (PolyML.makestring (length (fact_graph |> Graph.keys), length (fact_graph |> Graph.maximals),
blanchet@49333
   485
length (fact_graph |> Graph.minimals))) (*###*)
blanchet@49333
   486
    val parents = parents_wrt_facts facts fact_graph
blanchet@49333
   487
    val feats = features_of ctxt prover thy General (concl_t :: hyp_ts)
blanchet@49333
   488
    val suggs =
blanchet@49333
   489
      mash_QUERY ctxt overlord (max_suggs_of max_facts) (parents, feats)
blanchet@49326
   490
  in suggested_facts suggs facts end
blanchet@49264
   491
blanchet@49331
   492
fun add_thys_for thy =
blanchet@49331
   493
  let fun add comp thy = Symtab.update (Context.theory_name thy, comp) in
blanchet@49331
   494
    add false thy #> fold (add true) (Theory.ancestors_of thy)
blanchet@49331
   495
  end
blanchet@49319
   496
blanchet@49331
   497
fun update_fact_graph ctxt (name, parents, feats, deps) (upds, graph) =
blanchet@49331
   498
  let
blanchet@49331
   499
    fun maybe_add_from from (accum as (parents, graph)) =
blanchet@49331
   500
      (from :: parents, Graph.add_edge_acyclic (from, name) graph)
blanchet@49331
   501
      handle Graph.CYCLES _ =>
blanchet@49331
   502
             (trace_msg ctxt (fn () =>
blanchet@49331
   503
                  "Cycle between " ^ quote from ^ " and " ^ quote name); accum)
blanchet@49331
   504
           | Graph.UNDEF _ =>
blanchet@49331
   505
             (trace_msg ctxt (fn () => "Unknown node " ^ quote from); accum)
blanchet@49331
   506
    val graph = graph |> Graph.new_node (name, ())
blanchet@49331
   507
    val (parents, graph) = ([], graph) |> fold maybe_add_from parents
blanchet@49331
   508
    val (deps, graph) = ([], graph) |> fold maybe_add_from deps
blanchet@49331
   509
  in ((name, parents, feats, deps) :: upds, graph) end
blanchet@49321
   510
blanchet@49333
   511
val pass1_learn_timeout_factor = 0.5
blanchet@49333
   512
val pass2_learn_timeout_factor = 10.0
blanchet@49333
   513
blanchet@49333
   514
(* The timeout is understood in a very slack fashion. *)
blanchet@49333
   515
fun mash_learn_thy ctxt ({provers, verbose, overlord, ...} : params) thy
blanchet@49333
   516
                   timeout =
blanchet@49319
   517
  let
blanchet@49333
   518
    val timer = Timer.startRealTimer ()
blanchet@49333
   519
    val prover = hd provers
blanchet@49333
   520
    fun timed_out frac =
blanchet@49333
   521
      Time.> (Timer.checkRealTimer timer, time_mult frac timeout)
blanchet@49333
   522
    val css_table = clasimpset_rule_table_of ctxt
blanchet@49330
   523
    val facts = all_facts_of thy css_table
blanchet@49331
   524
    val {fact_graph, ...} = mash_get ()
blanchet@49331
   525
    fun is_old (_, th) = can (Graph.get_node fact_graph) (Thm.get_name_hint th)
blanchet@49331
   526
    val new_facts = facts |> filter_out is_old |> sort (thm_ord o pairself snd)
blanchet@49323
   527
  in
blanchet@49323
   528
    if null new_facts then
blanchet@49323
   529
      ()
blanchet@49323
   530
    else
blanchet@49319
   531
      let
blanchet@49333
   532
        val n = length new_facts
blanchet@49333
   533
        val _ =
blanchet@49333
   534
          if verbose then
blanchet@49333
   535
            "MaShing " ^ string_of_int n ^ " fact" ^ plural_s n ^
blanchet@49333
   536
            " (advisory timeout: " ^ string_from_time timeout ^ ")..."
blanchet@49333
   537
            |> Output.urgent_message
blanchet@49333
   538
          else
blanchet@49333
   539
            ()
blanchet@49323
   540
        val ths = facts |> map snd
blanchet@49330
   541
        val all_names =
blanchet@49333
   542
          ths |> filter_out (is_likely_tautology ctxt prover orf is_too_meta)
blanchet@49331
   543
              |> map (rpair () o Thm.get_name_hint)
blanchet@49331
   544
              |> Symtab.make
blanchet@49333
   545
        fun do_fact _ (accum as (_, true)) = accum
blanchet@49333
   546
          | do_fact ((_, (_, status)), th) ((parents, upds), false) =
blanchet@49333
   547
            let
blanchet@49333
   548
              val name = Thm.get_name_hint th
blanchet@49333
   549
              val feats = features_of ctxt prover thy status [prop_of th]
blanchet@49333
   550
              val deps = isabelle_dependencies_of all_names th
blanchet@49333
   551
              val upd = (name, parents, feats, deps)
blanchet@49333
   552
            in (([name], upd :: upds), timed_out pass1_learn_timeout_factor) end
blanchet@49333
   553
        val parents = parents_wrt_facts facts fact_graph
blanchet@49333
   554
        val ((_, upds), _) =
blanchet@49333
   555
          ((parents, []), false) |> fold do_fact new_facts |>> apsnd rev
blanchet@49333
   556
        val n = length upds
blanchet@49331
   557
        fun trans {thys, fact_graph} =
blanchet@49331
   558
          let
blanchet@49331
   559
            val mash_INIT_or_ADD =
blanchet@49331
   560
              if Graph.is_empty fact_graph then mash_INIT else mash_ADD
blanchet@49331
   561
            val (upds, fact_graph) =
blanchet@49331
   562
              ([], fact_graph) |> fold (update_fact_graph ctxt) upds
blanchet@49331
   563
          in
blanchet@49331
   564
            (mash_INIT_or_ADD ctxt overlord (rev upds);
blanchet@49331
   565
             {thys = thys |> add_thys_for thy,
blanchet@49331
   566
              fact_graph = fact_graph})
blanchet@49331
   567
          end
blanchet@49333
   568
      in
blanchet@49333
   569
        TimeLimit.timeLimit (time_mult pass2_learn_timeout_factor timeout)
blanchet@49333
   570
                            mash_map trans
blanchet@49333
   571
        handle TimeLimit.TimeOut =>
blanchet@49333
   572
               (if verbose then
blanchet@49333
   573
                  "MaSh timed out trying to learn " ^ string_of_int n ^
blanchet@49333
   574
                  " fact" ^ plural_s n ^ " in " ^
blanchet@49333
   575
                  string_from_time (Timer.checkRealTimer timer) ^ "."
blanchet@49333
   576
                  |> Output.urgent_message
blanchet@49333
   577
                else
blanchet@49333
   578
                  ());
blanchet@49333
   579
        (if verbose then
blanchet@49333
   580
           "MaSh learned " ^ string_of_int n ^ " fact" ^ plural_s n ^ " in " ^
blanchet@49333
   581
           string_from_time (Timer.checkRealTimer timer) ^ "."
blanchet@49333
   582
           |> Output.urgent_message
blanchet@49333
   583
         else
blanchet@49333
   584
           ())
blanchet@49333
   585
      end
blanchet@49323
   586
  end
blanchet@49319
   587
blanchet@49333
   588
fun mash_learn_proof ctxt ({provers, overlord, ...} : params) t used_ths facts =
blanchet@49331
   589
  let
blanchet@49331
   590
    val thy = Proof_Context.theory_of ctxt
blanchet@49333
   591
    val prover = hd provers
blanchet@49333
   592
    val name = ATP_Util.timestamp () ^ serial_string () (* fresh enough *)
blanchet@49333
   593
    val feats = features_of ctxt prover thy General [t]
blanchet@49331
   594
    val deps = used_ths |> map Thm.get_name_hint
blanchet@49331
   595
  in
blanchet@49331
   596
    mash_map (fn {thys, fact_graph} =>
blanchet@49324
   597
        let
blanchet@49333
   598
          val parents = parents_wrt_facts facts fact_graph
blanchet@49331
   599
          val upds = [(name, parents, feats, deps)]
blanchet@49331
   600
          val (upds, fact_graph) =
blanchet@49331
   601
            ([], fact_graph) |> fold (update_fact_graph ctxt) upds
blanchet@49324
   602
        in
blanchet@49331
   603
          mash_ADD ctxt overlord upds;
blanchet@49331
   604
          {thys = thys, fact_graph = fact_graph}
blanchet@49331
   605
        end)
blanchet@49331
   606
  end
blanchet@49264
   607
blanchet@49333
   608
(* The threshold should be large enough so that MaSh doesn't kick in for Auto
blanchet@49333
   609
   Sledgehammer and Try. *)
blanchet@49333
   610
val min_secs_for_learning = 15
blanchet@49333
   611
val short_learn_timeout_factor = 0.2
blanchet@49333
   612
val long_learn_timeout_factor = 4.0
blanchet@49333
   613
blanchet@49333
   614
fun relevant_facts ctxt (params as {fact_filter, timeout, ...}) prover max_facts
blanchet@49313
   615
        ({add, only, ...} : fact_override) hyp_ts concl_t facts =
blanchet@49329
   616
  if not (subset (op =) (the_list fact_filter, fact_filters)) then
blanchet@49329
   617
    error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".")
blanchet@49329
   618
  else if only then
blanchet@49304
   619
    facts
blanchet@49308
   620
  else if max_facts <= 0 then
blanchet@49303
   621
    []
blanchet@49303
   622
  else
blanchet@49303
   623
    let
blanchet@49333
   624
      val thy = Proof_Context.theory_of ctxt
blanchet@49333
   625
      fun maybe_learn can_suggest =
blanchet@49333
   626
        if Time.toSeconds timeout >= min_secs_for_learning then
blanchet@49333
   627
          if Multithreading.enabled () then
blanchet@49333
   628
            let
blanchet@49333
   629
              val factor =
blanchet@49333
   630
                if can_suggest then short_learn_timeout_factor
blanchet@49333
   631
                else long_learn_timeout_factor
blanchet@49333
   632
            in
blanchet@49333
   633
              Future.fork (fn () => mash_learn_thy ctxt params thy
blanchet@49333
   634
                                        (time_mult factor timeout)); ()
blanchet@49333
   635
            end
blanchet@49333
   636
          else
blanchet@49333
   637
            mash_learn_thy ctxt params thy
blanchet@49333
   638
                           (time_mult short_learn_timeout_factor timeout)
blanchet@49333
   639
        else
blanchet@49333
   640
          ()
blanchet@49329
   641
      val fact_filter =
blanchet@49329
   642
        case fact_filter of
blanchet@49333
   643
          SOME ff =>
blanchet@49333
   644
          (if ff <> iterN then maybe_learn (mash_can_suggest_facts ()) else ();
blanchet@49333
   645
           ff)
blanchet@49333
   646
        | NONE =>
blanchet@49333
   647
          if mash_can_suggest_facts () then (maybe_learn true; meshN)
blanchet@49333
   648
          else if mash_could_suggest_facts () then (maybe_learn false; iterN)
blanchet@49333
   649
          else iterN
blanchet@49303
   650
      val add_ths = Attrib.eval_thms ctxt add
blanchet@49307
   651
      fun prepend_facts ths accepts =
blanchet@49303
   652
        ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
blanchet@49307
   653
         (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
blanchet@49308
   654
        |> take max_facts
blanchet@49329
   655
      fun iter () =
blanchet@49313
   656
        iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts
blanchet@49313
   657
                                 concl_t facts
blanchet@49329
   658
        |> (fn facts => (facts, SOME (length facts)))
blanchet@49329
   659
      fun mash () =
blanchet@49333
   660
        (mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts,
blanchet@49333
   661
         NONE)
blanchet@49329
   662
      val mess =
blanchet@49329
   663
        [] |> (if fact_filter <> mashN then cons (iter ()) else I)
blanchet@49329
   664
           |> (if fact_filter <> iterN then cons (mash ()) else I)
blanchet@49303
   665
    in
blanchet@49328
   666
      mesh_facts max_facts mess
blanchet@49303
   667
      |> not (null add_ths) ? prepend_facts add_ths
blanchet@49303
   668
    end
blanchet@49303
   669
blanchet@49263
   670
end;