src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
author blanchet
Wed, 18 Jul 2012 08:44:03 +0200
changeset 49317 6cf5e58f1185
parent 49316 e5c5037a3104
child 49318 f1d135d0ea69
permissions -rw-r--r--
more implementation work on MaSh
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@49266
    17
  val fact_name_of : string -> string
blanchet@49314
    18
  val all_non_tautological_facts_of :
blanchet@49314
    19
    theory -> status Termtab.table -> fact list
blanchet@49266
    20
  val theory_ord : theory * theory -> order
blanchet@49266
    21
  val thm_ord : thm * thm -> order
blanchet@49266
    22
  val thms_by_thy : ('a * thm) list -> (string * thm list) list
blanchet@49266
    23
  val has_thy : theory -> thm -> bool
blanchet@49311
    24
  val parent_facts : (string * thm list) list -> theory -> string list
blanchet@49317
    25
  val features_of : theory -> status -> term list -> string list
blanchet@49266
    26
  val isabelle_dependencies_of : string list -> thm -> string list
blanchet@49266
    27
  val goal_of_thm : theory -> thm -> thm
blanchet@49311
    28
  val run_prover : Proof.context -> params -> fact list -> thm -> prover_result
blanchet@49314
    29
  val generate_accessibility : Proof.context -> theory -> bool -> string -> unit
blanchet@49314
    30
  val generate_features : Proof.context -> theory -> bool -> string -> unit
blanchet@49314
    31
  val generate_isa_dependencies :
blanchet@49314
    32
    Proof.context -> theory -> bool -> string -> unit
blanchet@49266
    33
  val generate_atp_dependencies :
blanchet@49266
    34
    Proof.context -> params -> theory -> bool -> string -> unit
blanchet@49317
    35
  val mash_RESET : unit -> unit
blanchet@49317
    36
  val mash_ADD : string -> string list -> string list -> string list -> unit
blanchet@49317
    37
  val mash_DEL : string list -> string list -> unit
blanchet@49317
    38
  val mash_SUGGEST : string list -> string list -> string list
blanchet@49313
    39
  val mash_reset : unit -> unit
blanchet@49313
    40
  val mash_can_suggest_facts : unit -> bool
blanchet@49313
    41
  val mash_suggest_facts :
blanchet@49317
    42
    Proof.context -> params -> string -> int -> term list -> term -> fact list
blanchet@49313
    43
    -> fact list * fact list
blanchet@49313
    44
  val mash_can_learn_thy : theory -> bool
blanchet@49313
    45
  val mash_learn_thy : theory -> real -> unit
blanchet@49313
    46
  val mash_learn_proof : theory -> term -> thm list -> unit
blanchet@49303
    47
  val relevant_facts :
blanchet@49307
    48
    Proof.context -> params -> string -> int -> fact_override -> term list
blanchet@49311
    49
    -> term -> fact list -> fact list
blanchet@49263
    50
end;
blanchet@49263
    51
blanchet@49263
    52
structure Sledgehammer_Filter_MaSh : SLEDGEHAMMER_FILTER_MASH =
blanchet@49263
    53
struct
blanchet@49264
    54
blanchet@49266
    55
open ATP_Util
blanchet@49266
    56
open ATP_Problem_Generate
blanchet@49266
    57
open Sledgehammer_Util
blanchet@49266
    58
open Sledgehammer_Fact
blanchet@49266
    59
open Sledgehammer_Filter_Iter
blanchet@49266
    60
open Sledgehammer_Provers
blanchet@49266
    61
blanchet@49316
    62
val mash_dir = "mash"
blanchet@49316
    63
val model_file = "model"
blanchet@49316
    64
val state_file = "state"
blanchet@49316
    65
blanchet@49317
    66
fun mk_path file =
blanchet@49317
    67
  getenv "ISABELLE_HOME_USER" ^ "/" ^ mash_dir ^ "/" ^ file
blanchet@49317
    68
  |> Path.explode
blanchet@49316
    69
blanchet@49317
    70
val fresh_fact_prefix = Long_Name.separator
blanchet@49266
    71
blanchet@49266
    72
(*** Isabelle helpers ***)
blanchet@49266
    73
blanchet@49266
    74
fun escape_meta_char c =
blanchet@49266
    75
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
blanchet@49266
    76
     c = #")" orelse c = #"," then
blanchet@49266
    77
    String.str c
blanchet@49266
    78
  else
blanchet@49266
    79
    (* fixed width, in case more digits follow *)
blanchet@49266
    80
    "\\" ^ stringN_of_int 3 (Char.ord c)
blanchet@49266
    81
blanchet@49266
    82
val escape_meta = String.translate escape_meta_char
blanchet@49266
    83
blanchet@49266
    84
val thy_prefix = "y_"
blanchet@49266
    85
blanchet@49266
    86
val fact_name_of = escape_meta
blanchet@49266
    87
val thy_name_of = prefix thy_prefix o escape_meta
blanchet@49266
    88
val const_name_of = prefix const_prefix o escape_meta
blanchet@49266
    89
val type_name_of = prefix type_const_prefix o escape_meta
blanchet@49266
    90
val class_name_of = prefix class_prefix o escape_meta
blanchet@49266
    91
blanchet@49266
    92
val thy_name_of_thm = theory_of_thm #> Context.theory_name
blanchet@49266
    93
blanchet@49266
    94
local
blanchet@49266
    95
blanchet@49266
    96
fun has_bool @{typ bool} = true
blanchet@49266
    97
  | has_bool (Type (_, Ts)) = exists has_bool Ts
blanchet@49266
    98
  | has_bool _ = false
blanchet@49266
    99
blanchet@49266
   100
fun has_fun (Type (@{type_name fun}, _)) = true
blanchet@49266
   101
  | has_fun (Type (_, Ts)) = exists has_fun Ts
blanchet@49266
   102
  | has_fun _ = false
blanchet@49266
   103
blanchet@49266
   104
val is_conn = member (op =)
blanchet@49266
   105
  [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
blanchet@49266
   106
   @{const_name HOL.implies}, @{const_name Not},
blanchet@49266
   107
   @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex},
blanchet@49266
   108
   @{const_name HOL.eq}]
blanchet@49266
   109
blanchet@49266
   110
val has_bool_arg_const =
blanchet@49266
   111
  exists_Const (fn (c, T) =>
blanchet@49266
   112
                   not (is_conn c) andalso exists has_bool (binder_types T))
blanchet@49266
   113
blanchet@49266
   114
fun higher_inst_const thy (c, T) =
blanchet@49266
   115
  case binder_types T of
blanchet@49266
   116
    [] => false
blanchet@49266
   117
  | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts
blanchet@49266
   118
blanchet@49266
   119
val binders = [@{const_name All}, @{const_name Ex}]
blanchet@49266
   120
blanchet@49266
   121
in
blanchet@49266
   122
blanchet@49266
   123
fun is_fo_term thy t =
blanchet@49266
   124
  let
blanchet@49266
   125
    val t =
blanchet@49266
   126
      t |> Envir.beta_eta_contract
blanchet@49266
   127
        |> transform_elim_prop
blanchet@49266
   128
        |> Object_Logic.atomize_term thy
blanchet@49266
   129
  in
blanchet@49266
   130
    Term.is_first_order binders t andalso
blanchet@49266
   131
    not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T
blanchet@49266
   132
                          | _ => false) t orelse
blanchet@49266
   133
         has_bool_arg_const t orelse exists_Const (higher_inst_const thy) t)
blanchet@49266
   134
  end
blanchet@49266
   135
blanchet@49266
   136
end
blanchet@49266
   137
blanchet@49317
   138
fun interesting_terms_types_and_classes term_max_depth type_max_depth ts =
blanchet@49266
   139
  let
blanchet@49266
   140
    val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49266
   141
    val bad_consts = atp_widely_irrelevant_consts
blanchet@49266
   142
    fun do_add_type (Type (s, Ts)) =
blanchet@49266
   143
        (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
blanchet@49266
   144
        #> fold do_add_type Ts
blanchet@49266
   145
      | do_add_type (TFree (_, S)) = union (op =) (map class_name_of S)
blanchet@49266
   146
      | do_add_type (TVar (_, S)) = union (op =) (map class_name_of S)
blanchet@49266
   147
    fun add_type T = type_max_depth >= 0 ? do_add_type T
blanchet@49266
   148
    fun mk_app s args =
blanchet@49266
   149
      if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
blanchet@49266
   150
      else s
blanchet@49266
   151
    fun patternify ~1 _ = ""
blanchet@49266
   152
      | patternify depth t =
blanchet@49266
   153
        case strip_comb t of
blanchet@49266
   154
          (Const (s, _), args) =>
blanchet@49266
   155
          mk_app (const_name_of s) (map (patternify (depth - 1)) args)
blanchet@49266
   156
        | _ => ""
blanchet@49266
   157
    fun add_term_patterns ~1 _ = I
blanchet@49266
   158
      | add_term_patterns depth t =
blanchet@49266
   159
        insert (op =) (patternify depth t)
blanchet@49266
   160
        #> add_term_patterns (depth - 1) t
blanchet@49266
   161
    val add_term = add_term_patterns term_max_depth
blanchet@49266
   162
    fun add_patterns t =
blanchet@49266
   163
      let val (head, args) = strip_comb t in
blanchet@49266
   164
        (case head of
blanchet@49266
   165
           Const (s, T) =>
blanchet@49266
   166
           not (member (op =) bad_consts s) ? (add_term t #> add_type T)
blanchet@49266
   167
         | Free (_, T) => add_type T
blanchet@49266
   168
         | Var (_, T) => add_type T
blanchet@49266
   169
         | Abs (_, T, body) => add_type T #> add_patterns body
blanchet@49266
   170
         | _ => I)
blanchet@49266
   171
        #> fold add_patterns args
blanchet@49266
   172
      end
blanchet@49317
   173
  in [] |> fold add_patterns ts |> sort string_ord end
blanchet@49266
   174
blanchet@49266
   175
fun is_likely_tautology th =
blanchet@49317
   176
  null (interesting_terms_types_and_classes 0 ~1 [prop_of th]) andalso
blanchet@49266
   177
  not (Thm.eq_thm_prop (@{thm ext}, th))
blanchet@49266
   178
blanchet@49266
   179
fun is_too_meta thy th =
blanchet@49266
   180
  fastype_of (Object_Logic.atomize_term thy (prop_of th)) <> @{typ bool}
blanchet@49266
   181
blanchet@49314
   182
fun all_non_tautological_facts_of thy css_table =
blanchet@49314
   183
  all_facts_of thy css_table
blanchet@49266
   184
  |> filter_out ((is_likely_tautology orf is_too_meta thy) o snd)
blanchet@49266
   185
blanchet@49266
   186
fun theory_ord p =
blanchet@49266
   187
  if Theory.eq_thy p then EQUAL
blanchet@49266
   188
  else if Theory.subthy p then LESS
blanchet@49266
   189
  else if Theory.subthy (swap p) then GREATER
blanchet@49266
   190
  else EQUAL
blanchet@49266
   191
blanchet@49266
   192
val thm_ord = theory_ord o pairself theory_of_thm
blanchet@49266
   193
blanchet@49266
   194
fun thms_by_thy ths =
blanchet@49266
   195
  ths |> map (snd #> `thy_name_of_thm)
blanchet@49266
   196
      |> AList.group (op =)
blanchet@49266
   197
      |> sort (int_ord o pairself (length o Theory.ancestors_of o theory_of_thm
blanchet@49266
   198
                                   o hd o snd))
blanchet@49266
   199
      |> map (apsnd (sort thm_ord))
blanchet@49266
   200
blanchet@49266
   201
fun has_thy thy th = (Context.theory_name thy = thy_name_of_thm th)
blanchet@49266
   202
blanchet@49311
   203
fun add_last_thms thy_ths thy =
blanchet@49311
   204
  case AList.lookup (op =) thy_ths (Context.theory_name thy) of
blanchet@49311
   205
    SOME (ths as _ :: _) => insert Thm.eq_thm (List.last ths)
blanchet@49311
   206
  | _ => add_parent_thms thy_ths thy
blanchet@49311
   207
and add_parent_thms thy_ths thy =
blanchet@49311
   208
  fold (add_last_thms thy_ths) (Theory.parents_of thy)
blanchet@49311
   209
blanchet@49311
   210
fun parent_facts thy_ths thy =
blanchet@49311
   211
  add_parent_thms thy_ths thy []
blanchet@49266
   212
  |> map (fact_name_of o Thm.get_name_hint)
blanchet@49266
   213
blanchet@49266
   214
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
blanchet@49266
   215
blanchet@49312
   216
val term_max_depth = 1
blanchet@49312
   217
val type_max_depth = 1
blanchet@49266
   218
blanchet@49266
   219
(* TODO: Generate type classes for types? *)
blanchet@49317
   220
fun features_of thy status ts =
blanchet@49317
   221
  thy_name_of (Context.theory_name thy) ::
blanchet@49317
   222
  interesting_terms_types_and_classes term_max_depth type_max_depth ts
blanchet@49317
   223
  |> exists (not o is_lambda_free) ts ? cons "lambdas"
blanchet@49317
   224
  |> exists (exists_Const is_exists) ts ? cons "skolems"
blanchet@49317
   225
  |> exists (not o is_fo_term thy) ts ? cons "ho"
blanchet@49317
   226
  |> (case status of
blanchet@49317
   227
        General => I
blanchet@49317
   228
      | Induction => cons "induction"
blanchet@49317
   229
      | Intro => cons "intro"
blanchet@49317
   230
      | Inductive => cons "inductive"
blanchet@49317
   231
      | Elim => cons "elim"
blanchet@49317
   232
      | Simp => cons "simp"
blanchet@49317
   233
      | Def => cons "def")
blanchet@49266
   234
blanchet@49266
   235
fun isabelle_dependencies_of all_facts =
blanchet@49266
   236
  thms_in_proof (SOME all_facts)
blanchet@49266
   237
  #> map fact_name_of #> sort string_ord
blanchet@49266
   238
blanchet@49266
   239
val freezeT = Type.legacy_freeze_type
blanchet@49266
   240
blanchet@49266
   241
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49266
   242
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49266
   243
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49266
   244
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49266
   245
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49266
   246
  | freeze t = t
blanchet@49266
   247
blanchet@49266
   248
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
blanchet@49266
   249
blanchet@49266
   250
fun run_prover ctxt (params as {provers, ...}) facts goal =
blanchet@49266
   251
  let
blanchet@49266
   252
    val problem =
blanchet@49266
   253
      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
blanchet@49304
   254
       facts = facts |> map (apfst (apfst (fn name => name ())))
blanchet@49304
   255
                     |> map Sledgehammer_Provers.Untranslated_Fact}
blanchet@49266
   256
    val prover =
blanchet@49266
   257
      Sledgehammer_Minimize.get_minimizing_prover ctxt
blanchet@49266
   258
          Sledgehammer_Provers.Normal (hd provers)
blanchet@49266
   259
  in prover params (K (K (K ""))) problem end
blanchet@49266
   260
blanchet@49314
   261
fun generate_accessibility ctxt thy include_thy file_name =
blanchet@49266
   262
  let
blanchet@49266
   263
    val path = file_name |> Path.explode
blanchet@49266
   264
    val _ = File.write path ""
blanchet@49314
   265
    val css_table = clasimpset_rule_table_of ctxt
blanchet@49266
   266
    fun do_thm th prevs =
blanchet@49266
   267
      let
blanchet@49266
   268
        val s = th ^ ": " ^ space_implode " " prevs ^ "\n"
blanchet@49266
   269
        val _ = File.append path s
blanchet@49266
   270
      in [th] end
blanchet@49266
   271
    val thy_ths =
blanchet@49314
   272
      all_non_tautological_facts_of thy css_table
blanchet@49266
   273
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49266
   274
      |> thms_by_thy
blanchet@49266
   275
    fun do_thy ths =
blanchet@49266
   276
      let
blanchet@49266
   277
        val thy = theory_of_thm (hd ths)
blanchet@49311
   278
        val parents = parent_facts thy_ths thy
blanchet@49266
   279
        val ths = ths |> map (fact_name_of o Thm.get_name_hint)
blanchet@49266
   280
      in fold do_thm ths parents; () end
blanchet@49266
   281
  in List.app (do_thy o snd) thy_ths end
blanchet@49266
   282
blanchet@49314
   283
fun generate_features ctxt thy include_thy file_name =
blanchet@49266
   284
  let
blanchet@49266
   285
    val path = file_name |> Path.explode
blanchet@49266
   286
    val _ = File.write path ""
blanchet@49314
   287
    val css_table = clasimpset_rule_table_of ctxt
blanchet@49266
   288
    val facts =
blanchet@49314
   289
      all_non_tautological_facts_of thy css_table
blanchet@49266
   290
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49266
   291
    fun do_fact ((_, (_, status)), th) =
blanchet@49266
   292
      let
blanchet@49266
   293
        val name = Thm.get_name_hint th
blanchet@49317
   294
        val feats = features_of (theory_of_thm th) status [prop_of th]
blanchet@49266
   295
        val s = fact_name_of name ^ ": " ^ space_implode " " feats ^ "\n"
blanchet@49266
   296
      in File.append path s end
blanchet@49266
   297
  in List.app do_fact facts end
blanchet@49266
   298
blanchet@49314
   299
fun generate_isa_dependencies ctxt thy include_thy file_name =
blanchet@49266
   300
  let
blanchet@49266
   301
    val path = file_name |> Path.explode
blanchet@49266
   302
    val _ = File.write path ""
blanchet@49314
   303
    val css_table = clasimpset_rule_table_of ctxt
blanchet@49266
   304
    val ths =
blanchet@49314
   305
      all_non_tautological_facts_of thy css_table
blanchet@49266
   306
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49266
   307
      |> map snd
blanchet@49266
   308
    val all_names = ths |> map Thm.get_name_hint
blanchet@49266
   309
    fun do_thm th =
blanchet@49266
   310
      let
blanchet@49266
   311
        val name = Thm.get_name_hint th
blanchet@49266
   312
        val deps = isabelle_dependencies_of all_names th
blanchet@49266
   313
        val s = fact_name_of name ^ ": " ^ space_implode " " deps ^ "\n"
blanchet@49266
   314
      in File.append path s end
blanchet@49266
   315
  in List.app do_thm ths end
blanchet@49266
   316
blanchet@49308
   317
fun generate_atp_dependencies ctxt (params as {provers, max_facts, ...}) thy
blanchet@49266
   318
                              include_thy file_name =
blanchet@49266
   319
  let
blanchet@49266
   320
    val path = file_name |> Path.explode
blanchet@49266
   321
    val _ = File.write path ""
blanchet@49314
   322
    val css_table = clasimpset_rule_table_of ctxt
blanchet@49266
   323
    val facts =
blanchet@49314
   324
      all_non_tautological_facts_of thy css_table
blanchet@49266
   325
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49266
   326
    val ths = facts |> map snd
blanchet@49266
   327
    val all_names = ths |> map Thm.get_name_hint
blanchet@49266
   328
    fun is_dep dep (_, th) = fact_name_of (Thm.get_name_hint th) = dep
blanchet@49266
   329
    fun add_isa_dep facts dep accum =
blanchet@49266
   330
      if exists (is_dep dep) accum then
blanchet@49266
   331
        accum
blanchet@49266
   332
      else case find_first (is_dep dep) facts of
blanchet@49304
   333
        SOME ((name, status), th) => accum @ [((name, status), th)]
blanchet@49266
   334
      | NONE => accum (* shouldn't happen *)
blanchet@49266
   335
    fun fix_name ((_, stature), th) =
blanchet@49304
   336
      ((fn () => th |> Thm.get_name_hint |> fact_name_of, stature), th)
blanchet@49266
   337
    fun do_thm th =
blanchet@49266
   338
      let
blanchet@49266
   339
        val name = Thm.get_name_hint th
blanchet@49266
   340
        val goal = goal_of_thm thy th
blanchet@49307
   341
        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
blanchet@49266
   342
        val deps =
blanchet@49266
   343
          case isabelle_dependencies_of all_names th of
blanchet@49266
   344
            [] => []
blanchet@49266
   345
          | isa_dep as [_] => isa_dep (* can hardly beat that *)
blanchet@49266
   346
          | isa_deps =>
blanchet@49266
   347
            let
blanchet@49266
   348
              val facts =
blanchet@49266
   349
                facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
blanchet@49266
   350
              val facts =
blanchet@49307
   351
                facts |> iterative_relevant_facts ctxt params (hd provers)
blanchet@49308
   352
                             (the max_facts) NONE hyp_ts concl_t
blanchet@49266
   353
                      |> fold (add_isa_dep facts) isa_deps
blanchet@49266
   354
                      |> map fix_name
blanchet@49266
   355
            in
blanchet@49266
   356
              case run_prover ctxt params facts goal of
blanchet@49266
   357
                {outcome = NONE, used_facts, ...} =>
blanchet@49266
   358
                used_facts |> map fst |> sort string_ord
blanchet@49266
   359
              | _ => isa_deps
blanchet@49266
   360
            end
blanchet@49266
   361
        val s = fact_name_of name ^ ": " ^ space_implode " " deps ^ "\n"
blanchet@49266
   362
      in File.append path s end
blanchet@49266
   363
  in List.app do_thm ths end
blanchet@49266
   364
blanchet@49266
   365
blanchet@49317
   366
(*** Low-level communication with MaSh ***)
blanchet@49317
   367
blanchet@49317
   368
fun mash_RESET () = (warning "MaSh RESET"; File.rm (mk_path model_file))
blanchet@49317
   369
blanchet@49317
   370
fun mash_ADD fact access feats deps =
blanchet@49317
   371
  warning ("MaSh ADD " ^ fact ^ ": " ^ space_implode " " access ^ "; " ^
blanchet@49317
   372
           space_implode " " feats ^ "; " ^ space_implode " " deps)
blanchet@49317
   373
blanchet@49317
   374
fun mash_DEL facts feats =
blanchet@49317
   375
  warning ("MaSh DEL " ^ space_implode " " facts ^ "; " ^
blanchet@49317
   376
           space_implode " " feats)
blanchet@49317
   377
blanchet@49317
   378
fun mash_SUGGEST access feats =
blanchet@49317
   379
  (warning ("MaSh SUGGEST " ^ space_implode " " access ^ "; " ^
blanchet@49317
   380
            space_implode " " feats);
blanchet@49317
   381
   [])
blanchet@49317
   382
blanchet@49317
   383
blanchet@49266
   384
(*** High-level communication with MaSh ***)
blanchet@49266
   385
blanchet@49316
   386
type mash_state =
blanchet@49317
   387
  {fresh : int,
blanchet@49317
   388
   completed_thys : unit Symtab.table,
blanchet@49316
   389
   thy_facts : string list Symtab.table}
blanchet@49264
   390
blanchet@49316
   391
val mash_zero =
blanchet@49317
   392
  {fresh = 0,
blanchet@49317
   393
   completed_thys = Symtab.empty,
blanchet@49316
   394
   thy_facts = Symtab.empty}
blanchet@49316
   395
blanchet@49316
   396
local
blanchet@49316
   397
blanchet@49316
   398
fun mash_load (state as (true, _)) = state
blanchet@49316
   399
  | mash_load _ =
blanchet@49317
   400
    let
blanchet@49317
   401
      val path = mk_path state_file
blanchet@49317
   402
      val _ = Isabelle_System.mkdir (path |> Path.dir)
blanchet@49317
   403
    in
blanchet@49317
   404
      (true,
blanchet@49317
   405
       case try File.read_lines path of
blanchet@49317
   406
         SOME (fresh_line :: comp_line :: facts_lines) =>
blanchet@49317
   407
         let
blanchet@49317
   408
           fun comp_thys_of_line comp_line =
blanchet@49317
   409
             Symtab.make (comp_line |> space_explode " " |> map (rpair ()))
blanchet@49317
   410
           fun add_facts_line line =
blanchet@49317
   411
             case space_explode " " line of
blanchet@49317
   412
               thy :: facts => Symtab.update_new (thy, facts)
blanchet@49317
   413
             | _ => I (* shouldn't happen *)
blanchet@49317
   414
         in
blanchet@49317
   415
           {fresh = Int.fromString fresh_line |> the_default 0,
blanchet@49317
   416
            completed_thys = comp_thys_of_line comp_line,
blanchet@49317
   417
            thy_facts = fold add_facts_line facts_lines Symtab.empty}
blanchet@49317
   418
         end
blanchet@49317
   419
       | _ => mash_zero)
blanchet@49317
   420
    end
blanchet@49316
   421
blanchet@49317
   422
fun mash_save ({fresh, completed_thys, thy_facts} : mash_state) =
blanchet@49316
   423
  let
blanchet@49317
   424
    val path = mk_path state_file
blanchet@49316
   425
    val comp_line = (completed_thys |> Symtab.keys |> space_implode " ") ^ "\n"
blanchet@49316
   426
    fun fact_line_for (thy, facts) = space_implode " " (thy :: facts) ^ "\n"
blanchet@49316
   427
  in
blanchet@49317
   428
    File.write path (string_of_int fresh ^ "\n" ^ comp_line);
blanchet@49316
   429
    Symtab.fold (fn thy_fact => fn () =>
blanchet@49316
   430
                    File.append path (fact_line_for thy_fact)) thy_facts
blanchet@49316
   431
  end
blanchet@49316
   432
blanchet@49317
   433
val global_state =
blanchet@49317
   434
  Synchronized.var "Sledgehammer_Filter_MaSh.global_state" (false, mash_zero)
blanchet@49316
   435
blanchet@49316
   436
in
blanchet@49316
   437
blanchet@49316
   438
fun mash_change f =
blanchet@49317
   439
  Synchronized.change global_state (mash_load ##> (f #> tap mash_save))
blanchet@49316
   440
blanchet@49317
   441
fun mash_peek f = Synchronized.change global_state (mash_load ##> tap f)
blanchet@49317
   442
blanchet@49317
   443
fun mash_value () = Synchronized.change_result global_state (mash_load #> `snd)
blanchet@49317
   444
blanchet@49317
   445
fun mash_reset () =
blanchet@49317
   446
  Synchronized.change global_state (fn _ =>
blanchet@49317
   447
      (mash_RESET (); File.rm (mk_path state_file); (true, mash_zero)))
blanchet@49316
   448
blanchet@49316
   449
end
blanchet@49316
   450
blanchet@49316
   451
fun mash_can_suggest_facts () =
blanchet@49316
   452
  not (Symtab.is_empty (#thy_facts (mash_value ())))
blanchet@49264
   453
blanchet@49317
   454
fun accessibility_of thy thy_facts =
blanchet@49317
   455
  let
blanchet@49317
   456
    val _ = 0
blanchet@49317
   457
  in
blanchet@49317
   458
    [] (*###*)
blanchet@49317
   459
  end
blanchet@49317
   460
blanchet@49313
   461
fun mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts =
blanchet@49316
   462
  let
blanchet@49317
   463
    val thy = Proof_Context.theory_of ctxt
blanchet@49317
   464
    val access = accessibility_of thy (#thy_facts (mash_value ()))
blanchet@49317
   465
    val feats = features_of thy General (concl_t :: hyp_ts)
blanchet@49316
   466
    val suggs = mash_SUGGEST access feats
blanchet@49316
   467
  in (facts, []) end
blanchet@49264
   468
blanchet@49316
   469
fun mash_can_learn_thy thy =
blanchet@49316
   470
  not (Symtab.defined (#completed_thys (mash_value ()))
blanchet@49316
   471
                      (Context.theory_name thy))
blanchet@49264
   472
blanchet@49313
   473
fun mash_learn_thy thy timeout = ()
blanchet@49317
   474
(* ### *)
blanchet@49313
   475
blanchet@49317
   476
fun mash_learn_proof thy t ths =
blanchet@49317
   477
  mash_change (fn {fresh, completed_thys, thy_facts} =>
blanchet@49317
   478
    let
blanchet@49317
   479
      val fact = fresh_fact_prefix ^ string_of_int fresh
blanchet@49317
   480
      val access = accessibility_of thy thy_facts
blanchet@49317
   481
      val feats = features_of thy General [t]
blanchet@49317
   482
      val deps = ths |> map (fact_name_of o Thm.get_name_hint)
blanchet@49317
   483
    in
blanchet@49317
   484
      mash_ADD fact access feats deps;
blanchet@49317
   485
      {fresh = fresh + 1, completed_thys = completed_thys,
blanchet@49317
   486
       thy_facts = thy_facts}
blanchet@49317
   487
    end)
blanchet@49264
   488
blanchet@49308
   489
fun relevant_facts ctxt params prover max_facts
blanchet@49313
   490
        ({add, only, ...} : fact_override) hyp_ts concl_t facts =
blanchet@49303
   491
  if only then
blanchet@49304
   492
    facts
blanchet@49308
   493
  else if max_facts <= 0 then
blanchet@49303
   494
    []
blanchet@49303
   495
  else
blanchet@49303
   496
    let
blanchet@49303
   497
      val add_ths = Attrib.eval_thms ctxt add
blanchet@49307
   498
      fun prepend_facts ths accepts =
blanchet@49303
   499
        ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
blanchet@49307
   500
         (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
blanchet@49308
   501
        |> take max_facts
blanchet@49313
   502
      val iter_facts =
blanchet@49313
   503
        iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts
blanchet@49313
   504
                                 concl_t facts
blanchet@49313
   505
      val (mash_facts, mash_rejected) =
blanchet@49313
   506
        mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts
blanchet@49317
   507
      val mesh_facts = iter_facts (* ### *)
blanchet@49303
   508
    in
blanchet@49313
   509
      mesh_facts
blanchet@49303
   510
      |> not (null add_ths) ? prepend_facts add_ths
blanchet@49303
   511
    end
blanchet@49303
   512
blanchet@49263
   513
end;