src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 49329 ee33ba3c0e05
parent 49328 0faafdffa662
child 49330 82d6e46c673f
permissions -rw-r--r--
added option to control which fact filter is used
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@49314
    29
  val all_non_tautological_facts_of :
blanchet@49314
    30
    theory -> status Termtab.table -> fact list
blanchet@49266
    31
  val theory_ord : theory * theory -> order
blanchet@49266
    32
  val thm_ord : thm * thm -> order
blanchet@49327
    33
  val thy_facts_from_thms : fact list -> string list Symtab.table
blanchet@49266
    34
  val has_thy : theory -> thm -> bool
blanchet@49318
    35
  val parent_facts : theory -> string list Symtab.table -> string list
blanchet@49317
    36
  val features_of : theory -> status -> term list -> string list
blanchet@49266
    37
  val isabelle_dependencies_of : string list -> thm -> string list
blanchet@49266
    38
  val goal_of_thm : theory -> thm -> thm
blanchet@49311
    39
  val run_prover : Proof.context -> params -> fact list -> thm -> prover_result
blanchet@49319
    40
  val thy_name_of_fact : string -> string
blanchet@49323
    41
  val mash_RESET : Proof.context -> unit
blanchet@49323
    42
  val mash_ADD :
blanchet@49323
    43
    Proof.context -> (string * string list * string list * string list) list
blanchet@49323
    44
    -> unit
blanchet@49323
    45
  val mash_DEL : Proof.context -> string list -> string list -> unit
blanchet@49326
    46
  val mash_QUERY : Proof.context -> string list * string list -> string list
blanchet@49323
    47
  val mash_reset : Proof.context -> unit
blanchet@49323
    48
  val mash_can_suggest_facts : Proof.context -> bool
blanchet@49313
    49
  val mash_suggest_facts :
blanchet@49326
    50
    Proof.context -> params -> string -> term list -> term -> fact list
blanchet@49326
    51
    -> fact list
blanchet@49323
    52
  val mash_can_learn_thy : Proof.context -> theory -> bool
blanchet@49323
    53
  val mash_learn_thy : Proof.context -> theory -> real -> unit
blanchet@49324
    54
  val mash_learn_proof : Proof.context -> theory -> term -> thm list -> unit
blanchet@49303
    55
  val relevant_facts :
blanchet@49307
    56
    Proof.context -> params -> string -> int -> fact_override -> term list
blanchet@49311
    57
    -> term -> fact list -> fact list
blanchet@49263
    58
end;
blanchet@49263
    59
blanchet@49263
    60
structure Sledgehammer_Filter_MaSh : SLEDGEHAMMER_FILTER_MASH =
blanchet@49263
    61
struct
blanchet@49264
    62
blanchet@49266
    63
open ATP_Util
blanchet@49266
    64
open ATP_Problem_Generate
blanchet@49266
    65
open Sledgehammer_Util
blanchet@49266
    66
open Sledgehammer_Fact
blanchet@49266
    67
open Sledgehammer_Filter_Iter
blanchet@49266
    68
open Sledgehammer_Provers
blanchet@49266
    69
blanchet@49323
    70
val trace =
blanchet@49323
    71
  Attrib.setup_config_bool @{binding sledgehammer_filter_mash_trace} (K false)
blanchet@49323
    72
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
blanchet@49323
    73
blanchet@49329
    74
val meshN = "mesh"
blanchet@49329
    75
val iterN = "iter"
blanchet@49329
    76
val mashN = "mash"
blanchet@49329
    77
blanchet@49329
    78
val fact_filters = [meshN, iterN, mashN]
blanchet@49329
    79
blanchet@49329
    80
fun mash_home () = getenv "MASH_HOME"
blanchet@49329
    81
fun mash_state_dir () =
blanchet@49324
    82
  getenv "ISABELLE_HOME_USER" ^ "/mash"
blanchet@49324
    83
  |> tap (fn dir => Isabelle_System.mkdir (Path.explode dir))
blanchet@49329
    84
fun mash_state_path () = mash_state_dir () ^ "/state" |> Path.explode
blanchet@49266
    85
blanchet@49266
    86
(*** Isabelle helpers ***)
blanchet@49266
    87
blanchet@49323
    88
fun meta_char c =
blanchet@49266
    89
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
blanchet@49266
    90
     c = #")" orelse c = #"," then
blanchet@49266
    91
    String.str c
blanchet@49266
    92
  else
blanchet@49266
    93
    (* fixed width, in case more digits follow *)
blanchet@49266
    94
    "\\" ^ stringN_of_int 3 (Char.ord c)
blanchet@49266
    95
blanchet@49323
    96
fun unmeta_chars accum [] = String.implode (rev accum)
blanchet@49323
    97
  | unmeta_chars accum (#"\\" :: d1 :: d2 :: d3 :: cs) =
blanchet@49323
    98
    (case Int.fromString (String.implode [d1, d2, d3]) of
blanchet@49323
    99
       SOME n => unmeta_chars (Char.chr n :: accum) cs
blanchet@49323
   100
     | NONE => "" (* error *))
blanchet@49323
   101
  | unmeta_chars _ (#"\\" :: _) = "" (* error *)
blanchet@49323
   102
  | unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs
blanchet@49323
   103
blanchet@49323
   104
val escape_meta = String.translate meta_char
blanchet@49318
   105
val escape_metas = map escape_meta #> space_implode " "
blanchet@49323
   106
val unescape_meta = unmeta_chars [] o String.explode
blanchet@49323
   107
val unescape_metas = map unescape_meta o space_explode " "
blanchet@49266
   108
blanchet@49326
   109
val explode_suggs =
blanchet@49326
   110
  space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta
blanchet@49326
   111
fun extract_query line =
blanchet@49326
   112
  case space_explode ":" line of
blanchet@49326
   113
    [goal_name, suggs] => (unescape_meta goal_name, explode_suggs suggs)
blanchet@49327
   114
  | _ => ("", [])
blanchet@49326
   115
blanchet@49326
   116
fun find_suggested facts sugg =
blanchet@49326
   117
  find_first (fn (_, th) => Thm.get_name_hint th = sugg) facts
blanchet@49326
   118
fun suggested_facts suggs facts = map_filter (find_suggested facts) suggs
blanchet@49326
   119
blanchet@49328
   120
fun sum_avg n xs =
blanchet@49328
   121
  fold (Integer.add o Integer.mult n) xs 0 div (length xs)
blanchet@49328
   122
blanchet@49329
   123
fun mesh_facts max_facts [(facts, _)] = facts |> take max_facts
blanchet@49329
   124
  | mesh_facts max_facts mess =
blanchet@49329
   125
    let
blanchet@49329
   126
      val n = length mess
blanchet@49329
   127
      val fact_eq = Thm.eq_thm o pairself snd
blanchet@49329
   128
      fun score_in fact (facts, def) =
blanchet@49329
   129
        case find_index (curry fact_eq fact) facts of
blanchet@49329
   130
          ~1 => def
blanchet@49329
   131
        | j => SOME j
blanchet@49329
   132
      fun score_of fact = mess |> map_filter (score_in fact) |> sum_avg n
blanchet@49329
   133
      val facts = fold (union fact_eq o take max_facts o fst) mess []
blanchet@49329
   134
    in
blanchet@49329
   135
      facts |> map (`score_of) |> sort (int_ord o pairself fst) |> map snd
blanchet@49329
   136
            |> take max_facts
blanchet@49329
   137
    end
blanchet@49327
   138
blanchet@49318
   139
val thy_feature_prefix = "y_"
blanchet@49266
   140
blanchet@49318
   141
val thy_feature_name_of = prefix thy_feature_prefix
blanchet@49318
   142
val const_name_of = prefix const_prefix
blanchet@49318
   143
val type_name_of = prefix type_const_prefix
blanchet@49318
   144
val class_name_of = prefix class_prefix
blanchet@49266
   145
blanchet@49266
   146
local
blanchet@49266
   147
blanchet@49266
   148
fun has_bool @{typ bool} = true
blanchet@49266
   149
  | has_bool (Type (_, Ts)) = exists has_bool Ts
blanchet@49266
   150
  | has_bool _ = false
blanchet@49266
   151
blanchet@49266
   152
fun has_fun (Type (@{type_name fun}, _)) = true
blanchet@49266
   153
  | has_fun (Type (_, Ts)) = exists has_fun Ts
blanchet@49266
   154
  | has_fun _ = false
blanchet@49266
   155
blanchet@49266
   156
val is_conn = member (op =)
blanchet@49266
   157
  [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
blanchet@49266
   158
   @{const_name HOL.implies}, @{const_name Not},
blanchet@49266
   159
   @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex},
blanchet@49266
   160
   @{const_name HOL.eq}]
blanchet@49266
   161
blanchet@49266
   162
val has_bool_arg_const =
blanchet@49266
   163
  exists_Const (fn (c, T) =>
blanchet@49266
   164
                   not (is_conn c) andalso exists has_bool (binder_types T))
blanchet@49266
   165
blanchet@49266
   166
fun higher_inst_const thy (c, T) =
blanchet@49266
   167
  case binder_types T of
blanchet@49266
   168
    [] => false
blanchet@49266
   169
  | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts
blanchet@49266
   170
blanchet@49266
   171
val binders = [@{const_name All}, @{const_name Ex}]
blanchet@49266
   172
blanchet@49266
   173
in
blanchet@49266
   174
blanchet@49266
   175
fun is_fo_term thy t =
blanchet@49266
   176
  let
blanchet@49266
   177
    val t =
blanchet@49266
   178
      t |> Envir.beta_eta_contract
blanchet@49266
   179
        |> transform_elim_prop
blanchet@49266
   180
        |> Object_Logic.atomize_term thy
blanchet@49266
   181
  in
blanchet@49266
   182
    Term.is_first_order binders t andalso
blanchet@49266
   183
    not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T
blanchet@49266
   184
                          | _ => false) t orelse
blanchet@49266
   185
         has_bool_arg_const t orelse exists_Const (higher_inst_const thy) t)
blanchet@49266
   186
  end
blanchet@49266
   187
blanchet@49266
   188
end
blanchet@49266
   189
blanchet@49317
   190
fun interesting_terms_types_and_classes term_max_depth type_max_depth ts =
blanchet@49266
   191
  let
blanchet@49266
   192
    val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49266
   193
    val bad_consts = atp_widely_irrelevant_consts
blanchet@49319
   194
    fun add_classes @{sort type} = I
blanchet@49319
   195
      | add_classes S = union (op =) (map class_name_of S)
blanchet@49266
   196
    fun do_add_type (Type (s, Ts)) =
blanchet@49266
   197
        (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
blanchet@49266
   198
        #> fold do_add_type Ts
blanchet@49319
   199
      | do_add_type (TFree (_, S)) = add_classes S
blanchet@49319
   200
      | do_add_type (TVar (_, S)) = add_classes S
blanchet@49266
   201
    fun add_type T = type_max_depth >= 0 ? do_add_type T
blanchet@49266
   202
    fun mk_app s args =
blanchet@49266
   203
      if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
blanchet@49266
   204
      else s
blanchet@49266
   205
    fun patternify ~1 _ = ""
blanchet@49266
   206
      | patternify depth t =
blanchet@49266
   207
        case strip_comb t of
blanchet@49266
   208
          (Const (s, _), args) =>
blanchet@49266
   209
          mk_app (const_name_of s) (map (patternify (depth - 1)) args)
blanchet@49266
   210
        | _ => ""
blanchet@49266
   211
    fun add_term_patterns ~1 _ = I
blanchet@49266
   212
      | add_term_patterns depth t =
blanchet@49266
   213
        insert (op =) (patternify depth t)
blanchet@49266
   214
        #> add_term_patterns (depth - 1) t
blanchet@49266
   215
    val add_term = add_term_patterns term_max_depth
blanchet@49266
   216
    fun add_patterns t =
blanchet@49266
   217
      let val (head, args) = strip_comb t in
blanchet@49266
   218
        (case head of
blanchet@49266
   219
           Const (s, T) =>
blanchet@49266
   220
           not (member (op =) bad_consts s) ? (add_term t #> add_type T)
blanchet@49266
   221
         | Free (_, T) => add_type T
blanchet@49266
   222
         | Var (_, T) => add_type T
blanchet@49266
   223
         | Abs (_, T, body) => add_type T #> add_patterns body
blanchet@49266
   224
         | _ => I)
blanchet@49266
   225
        #> fold add_patterns args
blanchet@49266
   226
      end
blanchet@49317
   227
  in [] |> fold add_patterns ts |> sort string_ord end
blanchet@49266
   228
blanchet@49266
   229
fun is_likely_tautology th =
blanchet@49317
   230
  null (interesting_terms_types_and_classes 0 ~1 [prop_of th]) andalso
blanchet@49266
   231
  not (Thm.eq_thm_prop (@{thm ext}, th))
blanchet@49266
   232
blanchet@49266
   233
fun is_too_meta thy th =
blanchet@49266
   234
  fastype_of (Object_Logic.atomize_term thy (prop_of th)) <> @{typ bool}
blanchet@49266
   235
blanchet@49314
   236
fun all_non_tautological_facts_of thy css_table =
blanchet@49314
   237
  all_facts_of thy css_table
blanchet@49266
   238
  |> filter_out ((is_likely_tautology orf is_too_meta thy) o snd)
blanchet@49266
   239
blanchet@49266
   240
fun theory_ord p =
blanchet@49266
   241
  if Theory.eq_thy p then EQUAL
blanchet@49266
   242
  else if Theory.subthy p then LESS
blanchet@49266
   243
  else if Theory.subthy (swap p) then GREATER
blanchet@49266
   244
  else EQUAL
blanchet@49266
   245
blanchet@49266
   246
val thm_ord = theory_ord o pairself theory_of_thm
blanchet@49266
   247
blanchet@49318
   248
(* ### FIXME: optimize *)
blanchet@49318
   249
fun thy_facts_from_thms ths =
blanchet@49318
   250
  ths |> map (snd #> `(theory_of_thm #> Context.theory_name))
blanchet@49266
   251
      |> AList.group (op =)
blanchet@49266
   252
      |> sort (int_ord o pairself (length o Theory.ancestors_of o theory_of_thm
blanchet@49266
   253
                                   o hd o snd))
blanchet@49318
   254
      |> map (apsnd (sort (rev_order o thm_ord) #> map Thm.get_name_hint))
blanchet@49318
   255
      |> Symtab.make
blanchet@49266
   256
blanchet@49318
   257
fun has_thy thy th =
blanchet@49318
   258
  Context.theory_name thy = Context.theory_name (theory_of_thm th)
blanchet@49266
   259
blanchet@49318
   260
fun parent_facts thy thy_facts =
blanchet@49318
   261
  let
blanchet@49318
   262
    fun add_last thy =
blanchet@49318
   263
      case Symtab.lookup thy_facts (Context.theory_name thy) of
blanchet@49318
   264
        SOME (last_fact :: _) => insert (op =) last_fact
blanchet@49318
   265
      | _ => add_parent thy
blanchet@49318
   266
    and add_parent thy = fold add_last (Theory.parents_of thy)
blanchet@49318
   267
  in add_parent thy [] end
blanchet@49266
   268
blanchet@49266
   269
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
blanchet@49266
   270
blanchet@49312
   271
val term_max_depth = 1
blanchet@49312
   272
val type_max_depth = 1
blanchet@49266
   273
blanchet@49266
   274
(* TODO: Generate type classes for types? *)
blanchet@49317
   275
fun features_of thy status ts =
blanchet@49318
   276
  thy_feature_name_of (Context.theory_name thy) ::
blanchet@49317
   277
  interesting_terms_types_and_classes term_max_depth type_max_depth ts
blanchet@49317
   278
  |> exists (not o is_lambda_free) ts ? cons "lambdas"
blanchet@49317
   279
  |> exists (exists_Const is_exists) ts ? cons "skolems"
blanchet@49317
   280
  |> exists (not o is_fo_term thy) ts ? cons "ho"
blanchet@49317
   281
  |> (case status of
blanchet@49317
   282
        General => I
blanchet@49317
   283
      | Induction => cons "induction"
blanchet@49317
   284
      | Intro => cons "intro"
blanchet@49317
   285
      | Inductive => cons "inductive"
blanchet@49317
   286
      | Elim => cons "elim"
blanchet@49317
   287
      | Simp => cons "simp"
blanchet@49317
   288
      | Def => cons "def")
blanchet@49266
   289
blanchet@49266
   290
fun isabelle_dependencies_of all_facts =
blanchet@49318
   291
  thms_in_proof (SOME all_facts) #> sort string_ord
blanchet@49266
   292
blanchet@49266
   293
val freezeT = Type.legacy_freeze_type
blanchet@49266
   294
blanchet@49266
   295
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49266
   296
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49266
   297
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49266
   298
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49266
   299
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49266
   300
  | freeze t = t
blanchet@49266
   301
blanchet@49266
   302
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
blanchet@49266
   303
blanchet@49266
   304
fun run_prover ctxt (params as {provers, ...}) facts goal =
blanchet@49266
   305
  let
blanchet@49266
   306
    val problem =
blanchet@49266
   307
      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
blanchet@49304
   308
       facts = facts |> map (apfst (apfst (fn name => name ())))
blanchet@49304
   309
                     |> map Sledgehammer_Provers.Untranslated_Fact}
blanchet@49266
   310
    val prover =
blanchet@49266
   311
      Sledgehammer_Minimize.get_minimizing_prover ctxt
blanchet@49266
   312
          Sledgehammer_Provers.Normal (hd provers)
blanchet@49266
   313
  in prover params (K (K (K ""))) problem end
blanchet@49266
   314
blanchet@49318
   315
fun accessibility_of thy thy_facts =
blanchet@49318
   316
  case Symtab.lookup thy_facts (Context.theory_name thy) of
blanchet@49318
   317
    SOME (fact :: _) => [fact]
blanchet@49318
   318
  | _ => parent_facts thy thy_facts
blanchet@49318
   319
blanchet@49319
   320
val thy_name_of_fact = hd o Long_Name.explode
blanchet@49266
   321
blanchet@49266
   322
blanchet@49317
   323
(*** Low-level communication with MaSh ***)
blanchet@49317
   324
blanchet@49326
   325
fun run_mash ctxt save write_cmds read_preds =
blanchet@49326
   326
  let
blanchet@49326
   327
    val temp_dir = getenv "ISABELLE_TMP"
blanchet@49326
   328
    val serial = serial_string ()
blanchet@49326
   329
    val cmd_file = temp_dir ^ "/mash_commands." ^ serial
blanchet@49326
   330
    val cmd_path = Path.explode cmd_file
blanchet@49326
   331
    val pred_file = temp_dir ^ "/mash_preds." ^ serial
blanchet@49326
   332
    val log_file = temp_dir ^ "/mash_log." ^ serial
blanchet@49326
   333
    val command =
blanchet@49329
   334
      mash_home () ^ "/mash.py --quiet --inputFile " ^ cmd_file ^
blanchet@49329
   335
      " --outputDir " ^ mash_state_dir () ^ " --predictions " ^ pred_file ^
blanchet@49328
   336
      " --log " ^ log_file ^ " --numberOfPredictions 1000" ^
blanchet@49329
   337
      (if save then " --saveModel" else "")
blanchet@49326
   338
    val _ = File.write cmd_path ""
blanchet@49326
   339
    val _ = write_cmds (File.append cmd_path)
blanchet@49326
   340
    val _ = trace_msg ctxt (fn () => "  running " ^ command)
blanchet@49326
   341
    val _ = Isabelle_System.bash command
blanchet@49326
   342
  in read_preds (fn () => File.read_lines (Path.explode pred_file)) end
blanchet@49326
   343
blanchet@49326
   344
fun str_of_update (fact, access, feats, deps) =
blanchet@49326
   345
  "! " ^ escape_meta fact ^ ": " ^ escape_metas access ^ "; " ^
blanchet@49326
   346
  escape_metas feats ^ "; " ^ escape_metas deps ^ "\n"
blanchet@49326
   347
blanchet@49326
   348
fun str_of_query (access, feats) =
blanchet@49326
   349
  "? " ^ escape_metas access ^ "; " ^ escape_metas feats
blanchet@49326
   350
blanchet@49323
   351
fun mash_RESET ctxt =
blanchet@49329
   352
  let val path = mash_state_dir () |> Path.explode in
blanchet@49324
   353
    trace_msg ctxt (K "MaSh RESET");
blanchet@49324
   354
    File.fold_dir (fn file => fn () =>
blanchet@49324
   355
                      File.rm (Path.append path (Path.basic file)))
blanchet@49324
   356
                  path ()
blanchet@49324
   357
  end
blanchet@49317
   358
blanchet@49324
   359
fun mash_ADD _ [] = ()
blanchet@49326
   360
  | mash_ADD ctxt upds =
blanchet@49326
   361
    (trace_msg ctxt (fn () => "MaSh ADD " ^ space_implode " " (map #1 upds));
blanchet@49326
   362
     run_mash ctxt true (fn append => List.app (append o str_of_update) upds)
blanchet@49326
   363
              (K ()))
blanchet@49317
   364
blanchet@49323
   365
fun mash_DEL ctxt facts feats =
blanchet@49323
   366
  trace_msg ctxt (fn () =>
blanchet@49323
   367
      "MaSh DEL " ^ escape_metas facts ^ "; " ^ escape_metas feats)
blanchet@49317
   368
blanchet@49326
   369
fun mash_QUERY ctxt (query as (_, feats)) =
blanchet@49329
   370
  (trace_msg ctxt (fn () => "MaSh QUERY " ^ space_implode " " feats);
blanchet@49326
   371
   run_mash ctxt false (fn append => append (str_of_query query))
blanchet@49326
   372
                 (fn preds => snd (extract_query (List.last (preds ()))))
blanchet@49326
   373
   handle List.Empty => [])
blanchet@49317
   374
blanchet@49317
   375
blanchet@49266
   376
(*** High-level communication with MaSh ***)
blanchet@49266
   377
blanchet@49316
   378
type mash_state =
blanchet@49323
   379
  {dirty_thys : unit Symtab.table,
blanchet@49316
   380
   thy_facts : string list Symtab.table}
blanchet@49264
   381
blanchet@49319
   382
val empty_state =
blanchet@49323
   383
  {dirty_thys = Symtab.empty,
blanchet@49316
   384
   thy_facts = Symtab.empty}
blanchet@49316
   385
blanchet@49316
   386
local
blanchet@49316
   387
blanchet@49316
   388
fun mash_load (state as (true, _)) = state
blanchet@49316
   389
  | mash_load _ =
blanchet@49324
   390
    let val path = mash_state_path () in
blanchet@49317
   391
      (true,
blanchet@49317
   392
       case try File.read_lines path of
blanchet@49323
   393
         SOME (dirty_line :: facts_lines) =>
blanchet@49317
   394
         let
blanchet@49321
   395
           fun dirty_thys_of_line line =
blanchet@49323
   396
             Symtab.make (line |> unescape_metas |> map (rpair ()))
blanchet@49317
   397
           fun add_facts_line line =
blanchet@49323
   398
             case unescape_metas line of
blanchet@49317
   399
               thy :: facts => Symtab.update_new (thy, facts)
blanchet@49317
   400
             | _ => I (* shouldn't happen *)
blanchet@49317
   401
         in
blanchet@49323
   402
           {dirty_thys = dirty_thys_of_line dirty_line,
blanchet@49317
   403
            thy_facts = fold add_facts_line facts_lines Symtab.empty}
blanchet@49317
   404
         end
blanchet@49319
   405
       | _ => empty_state)
blanchet@49317
   406
    end
blanchet@49316
   407
blanchet@49323
   408
fun mash_save ({dirty_thys, thy_facts} : mash_state) =
blanchet@49316
   409
  let
blanchet@49324
   410
    val path = mash_state_path ()
blanchet@49323
   411
    val dirty_line = (escape_metas (Symtab.keys dirty_thys)) ^ "\n"
blanchet@49318
   412
    fun fact_line_for (thy, facts) = escape_metas (thy :: facts) ^ "\n"
blanchet@49316
   413
  in
blanchet@49323
   414
    File.write path dirty_line;
blanchet@49316
   415
    Symtab.fold (fn thy_fact => fn () =>
blanchet@49323
   416
                    File.append path (fact_line_for thy_fact)) thy_facts ()
blanchet@49316
   417
  end
blanchet@49316
   418
blanchet@49317
   419
val global_state =
blanchet@49319
   420
  Synchronized.var "Sledgehammer_Filter_MaSh.global_state" (false, empty_state)
blanchet@49316
   421
blanchet@49316
   422
in
blanchet@49316
   423
blanchet@49321
   424
fun mash_map f =
blanchet@49317
   425
  Synchronized.change global_state (mash_load ##> (f #> tap mash_save))
blanchet@49316
   426
blanchet@49319
   427
fun mash_get () = Synchronized.change_result global_state (mash_load #> `snd)
blanchet@49317
   428
blanchet@49323
   429
fun mash_reset ctxt =
blanchet@49317
   430
  Synchronized.change global_state (fn _ =>
blanchet@49324
   431
      (mash_RESET ctxt; File.write (mash_state_path ()) "";
blanchet@49323
   432
       (true, empty_state)))
blanchet@49316
   433
blanchet@49316
   434
end
blanchet@49316
   435
blanchet@49323
   436
fun mash_can_suggest_facts (_ : Proof.context) =
blanchet@49323
   437
  not (Symtab.is_empty (#thy_facts (mash_get ())))
blanchet@49264
   438
blanchet@49326
   439
fun mash_suggest_facts ctxt params prover hyp_ts concl_t facts =
blanchet@49316
   440
  let
blanchet@49317
   441
    val thy = Proof_Context.theory_of ctxt
blanchet@49326
   442
    val thy_facts = #thy_facts (mash_get ())
blanchet@49326
   443
    val access = accessibility_of thy thy_facts
blanchet@49317
   444
    val feats = features_of thy General (concl_t :: hyp_ts)
blanchet@49326
   445
    val suggs = mash_QUERY ctxt (access, feats)
blanchet@49326
   446
  in suggested_facts suggs facts end
blanchet@49264
   447
blanchet@49323
   448
fun mash_can_learn_thy (_ : Proof.context) thy =
blanchet@49321
   449
  not (Symtab.defined (#dirty_thys (mash_get ())) (Context.theory_name thy))
blanchet@49264
   450
blanchet@49319
   451
fun is_fact_in_thy_facts thy_facts fact =
blanchet@49319
   452
  case Symtab.lookup thy_facts (thy_name_of_fact fact) of
blanchet@49319
   453
    SOME facts => member (op =) facts fact
blanchet@49319
   454
  | NONE => false
blanchet@49319
   455
blanchet@49321
   456
fun zip_facts news [] = news
blanchet@49321
   457
  | zip_facts [] olds = olds
blanchet@49321
   458
  | zip_facts (new :: news) (old :: olds) =
blanchet@49321
   459
    if new = old then
blanchet@49321
   460
      new :: zip_facts news olds
blanchet@49321
   461
    else if member (op =) news old then
blanchet@49321
   462
      old :: zip_facts (filter_out (curry (op =) old) news) olds
blanchet@49321
   463
    else if member (op =) olds new then
blanchet@49321
   464
      new :: zip_facts news (filter_out (curry (op =) new) olds)
blanchet@49321
   465
    else
blanchet@49321
   466
      new :: old :: zip_facts news olds
blanchet@49321
   467
blanchet@49321
   468
fun add_thy_facts_from_thys new old =
blanchet@49321
   469
  let
blanchet@49321
   470
    fun add_thy (thy, new_facts) =
blanchet@49321
   471
      case Symtab.lookup old thy of
blanchet@49321
   472
        SOME old_facts => Symtab.update (thy, zip_facts old_facts new_facts)
blanchet@49321
   473
      | NONE => Symtab.update_new (thy, new_facts)
blanchet@49321
   474
  in old |> Symtab.fold add_thy new end
blanchet@49321
   475
blanchet@49323
   476
fun mash_learn_thy ctxt thy timeout =
blanchet@49319
   477
  let
blanchet@49319
   478
    val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
blanchet@49319
   479
    val facts = all_non_tautological_facts_of thy css_table
blanchet@49321
   480
    val {thy_facts, ...} = mash_get ()
blanchet@49319
   481
    fun is_old (_, th) = is_fact_in_thy_facts thy_facts (Thm.get_name_hint th)
blanchet@49319
   482
    val (old_facts, new_facts) =
blanchet@49319
   483
      facts |> List.partition is_old ||> sort (thm_ord o pairself snd)
blanchet@49323
   484
  in
blanchet@49323
   485
    if null new_facts then
blanchet@49323
   486
      ()
blanchet@49323
   487
    else
blanchet@49319
   488
      let
blanchet@49323
   489
        val ths = facts |> map snd
blanchet@49323
   490
        val all_names = ths |> map Thm.get_name_hint
blanchet@49326
   491
        fun do_fact ((_, (_, status)), th) (prevs, upds) =
blanchet@49323
   492
          let
blanchet@49323
   493
            val name = Thm.get_name_hint th
blanchet@49323
   494
            val feats = features_of thy status [prop_of th]
blanchet@49323
   495
            val deps = isabelle_dependencies_of all_names th
blanchet@49326
   496
            val upd = (name, prevs, feats, deps)
blanchet@49326
   497
          in ([name], upd :: upds) end
blanchet@49323
   498
        val parents = parent_facts thy thy_facts
blanchet@49326
   499
        val (_, upds) = (parents, []) |> fold do_fact new_facts
blanchet@49323
   500
        val new_thy_facts = new_facts |> thy_facts_from_thms
blanchet@49323
   501
        fun trans {dirty_thys, thy_facts} =
blanchet@49326
   502
          (mash_ADD ctxt (rev upds);
blanchet@49323
   503
           {dirty_thys = dirty_thys,
blanchet@49323
   504
            thy_facts = thy_facts |> add_thy_facts_from_thys new_thy_facts})
blanchet@49323
   505
      in mash_map trans end
blanchet@49323
   506
  end
blanchet@49319
   507
blanchet@49324
   508
fun mash_learn_proof ctxt thy t ths =
blanchet@49324
   509
  mash_map (fn state as {dirty_thys, thy_facts} =>
blanchet@49324
   510
    let val deps = ths |> map Thm.get_name_hint in
blanchet@49324
   511
      if forall (is_fact_in_thy_facts thy_facts) deps then
blanchet@49324
   512
        let
blanchet@49324
   513
          val fact = ATP_Util.timestamp () (* should be fairly fresh *)
blanchet@49324
   514
          val access = accessibility_of thy thy_facts
blanchet@49324
   515
          val feats = features_of thy General [t]
blanchet@49324
   516
        in
blanchet@49324
   517
          mash_ADD ctxt [(fact, access, feats, deps)];
blanchet@49324
   518
          {dirty_thys = dirty_thys, thy_facts = thy_facts}
blanchet@49324
   519
        end
blanchet@49324
   520
      else
blanchet@49324
   521
        state
blanchet@49324
   522
    end)
blanchet@49264
   523
blanchet@49329
   524
fun relevant_facts ctxt (params as {fact_filter, ...}) prover max_facts
blanchet@49313
   525
        ({add, only, ...} : fact_override) hyp_ts concl_t facts =
blanchet@49329
   526
  if not (subset (op =) (the_list fact_filter, fact_filters)) then
blanchet@49329
   527
    error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".")
blanchet@49329
   528
  else if only then
blanchet@49304
   529
    facts
blanchet@49308
   530
  else if max_facts <= 0 then
blanchet@49303
   531
    []
blanchet@49303
   532
  else
blanchet@49303
   533
    let
blanchet@49329
   534
      val fact_filter =
blanchet@49329
   535
        case fact_filter of
blanchet@49329
   536
          SOME ff => ff
blanchet@49329
   537
        | NONE => if mash_home () = "" then iterN else meshN
blanchet@49303
   538
      val add_ths = Attrib.eval_thms ctxt add
blanchet@49307
   539
      fun prepend_facts ths accepts =
blanchet@49303
   540
        ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
blanchet@49307
   541
         (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
blanchet@49308
   542
        |> take max_facts
blanchet@49329
   543
      fun iter () =
blanchet@49313
   544
        iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts
blanchet@49313
   545
                                 concl_t facts
blanchet@49329
   546
        |> (fn facts => (facts, SOME (length facts)))
blanchet@49329
   547
      fun mash () =
blanchet@49329
   548
        (facts |> mash_suggest_facts ctxt params prover hyp_ts concl_t, NONE)
blanchet@49329
   549
      val mess =
blanchet@49329
   550
        [] |> (if fact_filter <> mashN then cons (iter ()) else I)
blanchet@49329
   551
           |> (if fact_filter <> iterN then cons (mash ()) else I)
blanchet@49303
   552
    in
blanchet@49328
   553
      mesh_facts max_facts mess
blanchet@49303
   554
      |> not (null add_ths) ? prepend_facts add_ths
blanchet@49303
   555
    end
blanchet@49303
   556
blanchet@49263
   557
end;