src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML
author blanchet
Thu, 26 Jun 2014 19:40:58 +0200
changeset 58728 5980ee29dbf6
parent 58727 0eb0c7b93676
child 58729 2b6fe2a48352
permissions -rw-r--r--
always expand all paths
blanchet@49395
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_mash.ML
blanchet@49263
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@58351
     3
    Author:     Cezary Kaliszyk, University of Innsbruck
blanchet@49263
     4
blanchet@49263
     5
Sledgehammer's machine-learning-based relevance filter (MaSh).
blanchet@49263
     6
*)
blanchet@49263
     7
blanchet@49396
     8
signature SLEDGEHAMMER_MASH =
blanchet@49263
     9
sig
blanchet@49266
    10
  type stature = ATP_Problem_Generate.stature
blanchet@52186
    11
  type raw_fact = Sledgehammer_Fact.raw_fact
blanchet@49311
    12
  type fact = Sledgehammer_Fact.fact
blanchet@49311
    13
  type fact_override = Sledgehammer_Fact.fact_override
blanchet@56543
    14
  type params = Sledgehammer_Prover.params
blanchet@56543
    15
  type prover_result = Sledgehammer_Prover.prover_result
blanchet@49266
    16
blanchet@49323
    17
  val trace : bool Config.T
blanchet@58492
    18
  val duplicates : bool Config.T
blanchet@52190
    19
  val MePoN : string
blanchet@49334
    20
  val MaShN : string
blanchet@52190
    21
  val MeShN : string
blanchet@49394
    22
  val mepoN : string
blanchet@49394
    23
  val mashN : string
blanchet@49329
    24
  val meshN : string
blanchet@49407
    25
  val unlearnN : string
blanchet@49407
    26
  val learn_isarN : string
blanchet@51499
    27
  val learn_proverN : string
blanchet@49407
    28
  val relearn_isarN : string
blanchet@51499
    29
  val relearn_proverN : string
blanchet@49329
    30
  val fact_filters : string list
blanchet@51841
    31
  val encode_str : string -> string
blanchet@51841
    32
  val encode_strs : string list -> string
blanchet@58347
    33
  val decode_str : string -> string
blanchet@58347
    34
  val decode_strs : string -> string list
blanchet@58397
    35
  val encode_features : (string * real) list -> string
blanchet@58467
    36
  val extract_suggestions : string -> string * (string * real) list
blanchet@51647
    37
blanchet@58620
    38
  datatype mash_engine =
blanchet@58620
    39
    MaSh_Py
blanchet@58620
    40
  | MaSh_SML_kNN
blanchet@58717
    41
  | MaSh_SML_kNN_Ext
blanchet@58704
    42
  | MaSh_SML_NB
blanchet@58717
    43
  | MaSh_SML_NB_Ext
blanchet@58448
    44
blanchet@58462
    45
  val is_mash_enabled : unit -> bool
blanchet@58462
    46
  val the_mash_engine : unit -> mash_engine
blanchet@58462
    47
blanchet@54285
    48
  val mash_unlearn : Proof.context -> params -> unit
blanchet@51639
    49
  val nickname_of_thm : thm -> string
blanchet@58348
    50
  val find_suggested_facts : Proof.context -> ('b * thm) list -> string list -> ('b * thm) list
blanchet@58348
    51
  val mesh_facts : ('a * 'a -> bool) -> int -> (real * (('a * real) list * 'a list)) list -> 'a list
blanchet@52272
    52
  val crude_thm_ord : thm * thm -> order
blanchet@52319
    53
  val thm_less : thm * thm -> bool
blanchet@49266
    54
  val goal_of_thm : theory -> thm -> thm
blanchet@58348
    55
  val run_prover_for_mash : Proof.context -> params -> string -> string -> fact list -> thm ->
blanchet@58348
    56
    prover_result
blanchet@58397
    57
  val features_of : Proof.context -> theory -> int -> int Symtab.table -> stature -> term list ->
blanchet@58397
    58
    (string * real) list
blanchet@52314
    59
  val trim_dependencies : string list -> string list option
blanchet@58648
    60
  val isar_dependencies_of : string Symtab.table * string Symtab.table -> thm -> string list option
blanchet@58348
    61
  val prover_dependencies_of : Proof.context -> params -> string -> int -> raw_fact list ->
blanchet@58348
    62
    string Symtab.table * string Symtab.table -> thm -> bool * string list
blanchet@58348
    63
  val attach_parents_to_facts : ('a * thm) list -> ('a * thm) list ->
blanchet@58348
    64
    (string list * ('a * thm)) list
blanchet@54277
    65
  val num_extra_feature_facts : int
blanchet@54278
    66
  val extra_feature_factor : real
blanchet@54277
    67
  val weight_facts_smoothly : 'a list -> ('a * real) list
blanchet@54277
    68
  val weight_facts_steeply : 'a list -> ('a * real) list
blanchet@58348
    69
  val find_mash_suggestions : Proof.context -> int -> string list -> ('b * thm) list ->
blanchet@58348
    70
    ('b * thm) list -> ('b * thm) list -> ('b * thm) list * ('b * thm) list
blanchet@54264
    71
  val add_const_counts : term -> int Symtab.table -> int Symtab.table
blanchet@58348
    72
  val mash_suggested_facts : Proof.context -> params -> int -> term list -> term -> raw_fact list ->
blanchet@58348
    73
    fact list * fact list
blanchet@55876
    74
  val mash_learn_proof : Proof.context -> params -> term -> ('a * thm) list -> thm list -> unit
blanchet@58462
    75
  val mash_learn_facts : Proof.context -> params -> string -> bool -> int -> bool -> Time.time ->
blanchet@58462
    76
    raw_fact list -> string
blanchet@58348
    77
  val mash_learn : Proof.context -> params -> fact_override -> thm list -> bool -> unit
blanchet@54956
    78
blanchet@54285
    79
  val mash_can_suggest_facts : Proof.context -> bool -> bool
blanchet@58450
    80
  val generous_max_suggestions : int -> int
blanchet@51829
    81
  val mepo_weight : real
blanchet@51829
    82
  val mash_weight : real
blanchet@58348
    83
  val relevant_facts : Proof.context -> params -> string -> int -> fact_override -> term list ->
blanchet@58348
    84
    term -> raw_fact list -> (string * fact list) list
blanchet@54285
    85
  val kill_learners : Proof.context -> params -> unit
blanchet@49334
    86
  val running_learners : unit -> unit
blanchet@49263
    87
end;
blanchet@49263
    88
blanchet@49396
    89
structure Sledgehammer_MaSh : SLEDGEHAMMER_MASH =
blanchet@49263
    90
struct
blanchet@49264
    91
blanchet@49266
    92
open ATP_Util
blanchet@49266
    93
open ATP_Problem_Generate
blanchet@49266
    94
open Sledgehammer_Util
blanchet@49266
    95
open Sledgehammer_Fact
blanchet@56543
    96
open Sledgehammer_Prover
blanchet@56544
    97
open Sledgehammer_Prover_Minimize
blanchet@49396
    98
open Sledgehammer_MePo
blanchet@49266
    99
blanchet@58347
   100
val trace = Attrib.setup_config_bool @{binding sledgehammer_mash_trace} (K false)
blanchet@58492
   101
val duplicates = Attrib.setup_config_bool @{binding sledgehammer_fact_duplicates} (K false)
blanchet@52214
   102
blanchet@49323
   103
fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
blanchet@49323
   104
blanchet@58492
   105
fun gen_eq_thm ctxt = if Config.get ctxt duplicates then Thm.eq_thm_strict else Thm.eq_thm_prop
blanchet@58492
   106
blanchet@52190
   107
val MePoN = "MePo"
blanchet@49334
   108
val MaShN = "MaSh"
blanchet@52190
   109
val MeShN = "MeSh"
blanchet@49334
   110
blanchet@49394
   111
val mepoN = "mepo"
blanchet@49394
   112
val mashN = "mash"
blanchet@49329
   113
val meshN = "mesh"
blanchet@49329
   114
blanchet@49394
   115
val fact_filters = [meshN, mepoN, mashN]
blanchet@49329
   116
blanchet@49407
   117
val unlearnN = "unlearn"
blanchet@49407
   118
val learn_isarN = "learn_isar"
blanchet@51499
   119
val learn_proverN = "learn_prover"
blanchet@49407
   120
val relearn_isarN = "relearn_isar"
blanchet@51499
   121
val relearn_proverN = "relearn_prover"
blanchet@49407
   122
blanchet@58710
   123
fun map_array_at ary f i = Array.update (ary, i, f (Array.sub (ary, i)))
blanchet@58710
   124
blanchet@58713
   125
type xtab = int * int Symtab.table
blanchet@58713
   126
blanchet@58713
   127
val empty_xtab = (0, Symtab.empty)
blanchet@58713
   128
blanchet@58713
   129
fun add_to_xtab key (next, tab) = (next + 1, Symtab.update_new (key, next) tab)
blanchet@58713
   130
fun maybe_add_to_xtab key = perhaps (try (add_to_xtab key))
blanchet@58713
   131
blanchet@58728
   132
fun mash_state_dir () =
blanchet@58728
   133
  Path.expand (Path.explode "$ISABELLE_HOME_USER/mash" |> tap Isabelle_System.mkdir)
blanchet@58728
   134
blanchet@51325
   135
fun mash_state_file () = Path.append (mash_state_dir ()) (Path.explode "state")
blanchet@49266
   136
blanchet@58349
   137
fun wipe_out_mash_state_dir () =
blanchet@58349
   138
  let val path = mash_state_dir () in
blanchet@58349
   139
    try (File.fold_dir (fn file => fn _ => try File.rm (Path.append path (Path.basic file))) path)
blanchet@58349
   140
      NONE;
blanchet@58349
   141
    ()
blanchet@58349
   142
  end
blanchet@49345
   143
blanchet@58620
   144
datatype mash_engine =
blanchet@58620
   145
  MaSh_Py
blanchet@58620
   146
| MaSh_SML_kNN
blanchet@58717
   147
| MaSh_SML_kNN_Ext
blanchet@58704
   148
| MaSh_SML_NB
blanchet@58717
   149
| MaSh_SML_NB_Ext
blanchet@58620
   150
blanchet@58370
   151
fun mash_engine () =
blanchet@58431
   152
  let val flag1 = Options.default_string @{system_option MaSh} in
blanchet@58370
   153
    (case if flag1 <> "none" (* default *) then flag1 else getenv "MASH" of
blanchet@58704
   154
      "yes" => SOME MaSh_SML_NB
blanchet@58370
   155
    | "py" => SOME MaSh_Py
blanchet@58704
   156
    | "sml" => SOME MaSh_SML_NB
blanchet@58371
   157
    | "sml_knn" => SOME MaSh_SML_kNN
blanchet@58717
   158
    | "sml_knn_ext" => SOME MaSh_SML_kNN_Ext
blanchet@58704
   159
    | "sml_nb" => SOME MaSh_SML_NB
blanchet@58717
   160
    | "sml_nb_ext" => SOME MaSh_SML_NB_Ext
blanchet@58370
   161
    | _ => NONE)
blanchet@58370
   162
  end
blanchet@58360
   163
blanchet@58370
   164
val is_mash_enabled = is_some o mash_engine
blanchet@58704
   165
val the_mash_engine = the_default MaSh_SML_NB o mash_engine
blanchet@58360
   166
blanchet@58349
   167
blanchet@58449
   168
(*** Low-level communication with the Python version of MaSh ***)
blanchet@51326
   169
blanchet@54254
   170
val save_models_arg = "--saveModels"
blanchet@54254
   171
val shutdown_server_arg = "--shutdownServer"
blanchet@54254
   172
blanchet@58464
   173
fun wipe_out_file file = ignore (try (File.rm o Path.explode) file)
blanchet@51326
   174
blanchet@51350
   175
fun write_file banner (xs, f) path =
blanchet@51350
   176
  (case banner of SOME s => File.write path s | NONE => ();
blanchet@54231
   177
   xs |> chunk_list 500 |> List.app (File.append path o implode o map f))
blanchet@51334
   178
  handle IO.Io _ => ()
blanchet@51326
   179
blanchet@54289
   180
fun run_mash_tool ctxt overlord extra_args background write_cmds read_suggs =
blanchet@51326
   181
  let
blanchet@51326
   182
    val (temp_dir, serial) =
blanchet@51326
   183
      if overlord then (getenv "ISABELLE_HOME_USER", "")
blanchet@51326
   184
      else (getenv "ISABELLE_TMP", serial_string ())
blanchet@54266
   185
    val log_file = temp_dir ^ "/mash_log" ^ serial
blanchet@51326
   186
    val err_file = temp_dir ^ "/mash_err" ^ serial
blanchet@51326
   187
    val sugg_file = temp_dir ^ "/mash_suggs" ^ serial
blanchet@51350
   188
    val sugg_path = Path.explode sugg_file
blanchet@51326
   189
    val cmd_file = temp_dir ^ "/mash_commands" ^ serial
blanchet@51350
   190
    val cmd_path = Path.explode cmd_file
blanchet@58349
   191
    val model_dir = File.shell_path (mash_state_dir ())
blanchet@58472
   192
blanchet@51326
   193
    val command =
blanchet@54693
   194
      "cd \"$ISABELLE_SLEDGEHAMMER_MASH\"/src; \
blanchet@54927
   195
      \PYTHONDONTWRITEBYTECODE=y ./mash.py\
blanchet@54927
   196
      \ --quiet\
blanchet@54927
   197
      \ --port=$MASH_PORT\
blanchet@54693
   198
      \ --outputDir " ^ model_dir ^
blanchet@54693
   199
      " --modelFile=" ^ model_dir ^ "/model.pickle\
blanchet@54693
   200
      \ --dictsFile=" ^ model_dir ^ "/dict.pickle\
blanchet@54927
   201
      \ --log " ^ log_file ^
blanchet@54927
   202
      " --inputFile " ^ cmd_file ^
blanchet@54927
   203
      " --predictions " ^ sugg_file ^
blanchet@58347
   204
      (if extra_args = [] then "" else " " ^ space_implode " " extra_args) ^ " >& " ^ err_file ^
blanchet@54289
   205
      (if background then " &" else "")
blanchet@58472
   206
blanchet@51326
   207
    fun run_on () =
blanchet@51765
   208
      (Isabelle_System.bash command
blanchet@55552
   209
       |> tap (fn _ =>
blanchet@58348
   210
         (case try File.read (Path.explode err_file) |> the_default "" of
blanchet@58348
   211
           "" => trace_msg ctxt (K "Done")
blanchet@58348
   212
         | s => warning ("MaSh error: " ^ elide_string 1000 s)));
blanchet@51350
   213
       read_suggs (fn () => try File.read_lines sugg_path |> these))
blanchet@58472
   214
blanchet@51326
   215
    fun clean_up () =
blanchet@58347
   216
      if overlord then () else List.app wipe_out_file [err_file, sugg_file, cmd_file]
blanchet@51326
   217
  in
blanchet@51350
   218
    write_file (SOME "") ([], K "") sugg_path;
blanchet@51350
   219
    write_file (SOME "") write_cmds cmd_path;
blanchet@51326
   220
    trace_msg ctxt (fn () => "Running " ^ command);
blanchet@51326
   221
    with_cleanup clean_up run_on ()
blanchet@51326
   222
  end
blanchet@49266
   223
blanchet@49323
   224
fun meta_char c =
blanchet@58348
   225
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse c = #")" orelse
blanchet@58348
   226
     c = #"," then
blanchet@49266
   227
    String.str c
blanchet@49266
   228
  else
blanchet@49266
   229
    (* fixed width, in case more digits follow *)
blanchet@49410
   230
    "%" ^ stringN_of_int 3 (Char.ord c)
blanchet@49266
   231
blanchet@49323
   232
fun unmeta_chars accum [] = String.implode (rev accum)
blanchet@49410
   233
  | unmeta_chars accum (#"%" :: d1 :: d2 :: d3 :: cs) =
blanchet@49323
   234
    (case Int.fromString (String.implode [d1, d2, d3]) of
blanchet@58347
   235
      SOME n => unmeta_chars (Char.chr n :: accum) cs
blanchet@58347
   236
    | NONE => "" (* error *))
blanchet@49410
   237
  | unmeta_chars _ (#"%" :: _) = "" (* error *)
blanchet@49323
   238
  | unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs
blanchet@49323
   239
blanchet@51841
   240
val encode_str = String.translate meta_char
blanchet@58347
   241
val decode_str = String.explode #> unmeta_chars []
blanchet@58347
   242
blanchet@51841
   243
val encode_strs = map encode_str #> space_implode " "
blanchet@58347
   244
val decode_strs = space_explode " " #> filter_out (curry (op =) "") #> map decode_str
blanchet@49266
   245
blanchet@54695
   246
(* Avoid scientific notation *)
blanchet@54695
   247
fun safe_str_of_real r =
blanchet@54695
   248
  if r < 0.00001 then "0.00001"
blanchet@54695
   249
  else if r >= 1000000.0 then "1000000"
blanchet@54695
   250
  else Markup.print_real r
blanchet@54695
   251
blanchet@56037
   252
fun encode_feature (names, weight) =
blanchet@58397
   253
  encode_str names ^ (if Real.== (weight, 1.0) then "" else "=" ^ safe_str_of_real weight)
blanchet@56037
   254
blanchet@51371
   255
val encode_features = map encode_feature #> space_implode " "
blanchet@51371
   256
blanchet@58347
   257
fun str_of_learn (name, parents, feats, deps) =
blanchet@58397
   258
  "! " ^ encode_str name ^ ": " ^ encode_strs parents ^ "; " ^ encode_strs feats ^ "; " ^
blanchet@58397
   259
  encode_strs deps ^ "\n"
blanchet@49684
   260
blanchet@58347
   261
fun str_of_relearn (name, deps) = "p " ^ encode_str name ^ ": " ^ encode_strs deps ^ "\n"
blanchet@49684
   262
blanchet@54236
   263
fun str_of_query max_suggs (learns, hints, parents, feats) =
blanchet@54232
   264
  implode (map str_of_learn learns) ^
blanchet@58347
   265
  "? " ^ string_of_int max_suggs ^ " # " ^ encode_strs parents ^ "; " ^ encode_features feats ^
blanchet@54232
   266
  (if null hints then "" else "; " ^ encode_strs hints) ^ "\n"
blanchet@49421
   267
blanchet@58347
   268
(* The suggested weights do not make much sense. *)
blanchet@49421
   269
fun extract_suggestion sugg =
blanchet@56628
   270
  (case space_explode "=" sugg of
blanchet@58467
   271
    [name, weight] => SOME (decode_str name, Real.fromString weight |> the_default 1.0)
blanchet@58467
   272
  | [name] => SOME (decode_str name, 1.0)
blanchet@56628
   273
  | _ => NONE)
blanchet@49421
   274
blanchet@51648
   275
fun extract_suggestions line =
blanchet@56628
   276
  (case space_explode ":" line of
blanchet@58347
   277
    [goal, suggs] => (decode_str goal, map_filter extract_suggestion (space_explode " " suggs))
blanchet@56628
   278
  | _ => ("", []))
blanchet@49326
   279
blanchet@58349
   280
structure MaSh_Py =
blanchet@51647
   281
struct
blanchet@51647
   282
blanchet@54285
   283
fun shutdown ctxt overlord =
blanchet@58349
   284
  (trace_msg ctxt (K "MaSh_Py shutdown");
blanchet@54893
   285
   run_mash_tool ctxt overlord [shutdown_server_arg] false ([], K "") (K ()))
blanchet@54289
   286
blanchet@54289
   287
fun save ctxt overlord =
blanchet@58349
   288
  (trace_msg ctxt (K "MaSh_Py save");
blanchet@54290
   289
   run_mash_tool ctxt overlord [save_models_arg] true ([], K "") (K ()))
blanchet@54279
   290
blanchet@54285
   291
fun unlearn ctxt overlord =
blanchet@58349
   292
  (trace_msg ctxt (K "MaSh_Py unlearn");
blanchet@58349
   293
   shutdown ctxt overlord;
blanchet@58349
   294
   wipe_out_mash_state_dir ())
blanchet@51326
   295
blanchet@54894
   296
fun learn _ _ _ [] = ()
blanchet@58348
   297
  | learn ctxt overlord save learns =
blanchet@58349
   298
    (trace_msg ctxt (fn () =>
blanchet@58351
   299
       "MaSh_Py learn {" ^ elide_string 1000 (space_implode " " (map #1 learns)) ^ "}");
blanchet@58349
   300
     run_mash_tool ctxt overlord ([] |> save ? cons save_models_arg) false (learns, str_of_learn)
blanchet@58349
   301
       (K ()))
blanchet@51326
   302
blanchet@54894
   303
fun relearn _ _ _ [] = ()
blanchet@54894
   304
  | relearn ctxt overlord save relearns =
blanchet@58349
   305
    (trace_msg ctxt (fn () => "MaSh_Py relearn " ^
blanchet@58348
   306
       elide_string 1000 (space_implode " " (map #1 relearns)));
blanchet@54894
   307
     run_mash_tool ctxt overlord ([] |> save ? cons save_models_arg) false
blanchet@58348
   308
       (relearns, str_of_relearn) (K ()))
blanchet@51326
   309
blanchet@54254
   310
fun query ctxt overlord max_suggs (query as (_, _, _, feats)) =
blanchet@58349
   311
  (trace_msg ctxt (fn () => "MaSh_Py query " ^ encode_features feats);
blanchet@56628
   312
   run_mash_tool ctxt overlord [] false ([query], str_of_query max_suggs) (fn suggs =>
blanchet@58349
   313
     (case suggs () of [] => [] | suggs => snd (extract_suggestions (List.last suggs))))
blanchet@51326
   314
   handle List.Empty => [])
blanchet@51326
   315
blanchet@51647
   316
end;
blanchet@51647
   317
blanchet@51326
   318
blanchet@58349
   319
(*** Standard ML version of MaSh ***)
blanchet@58349
   320
blanchet@58349
   321
structure MaSh_SML =
blanchet@58349
   322
struct
blanchet@58349
   323
blanchet@58351
   324
exception BOTTOM of int
blanchet@58351
   325
blanchet@58697
   326
fun heap cmp bnd al a =
blanchet@58351
   327
  let
blanchet@58351
   328
    fun maxson l i =
blanchet@58359
   329
      let val i31 = i + i + i + 1 in
blanchet@58351
   330
        if i31 + 2 < l then
blanchet@58359
   331
          let val x = Unsynchronized.ref i31 in
blanchet@58359
   332
            if cmp (Array.sub (a, i31), Array.sub (a, i31 + 1)) = LESS then x := i31 + 1 else ();
blanchet@58359
   333
            if cmp (Array.sub (a, !x), Array.sub (a, i31 + 2)) = LESS then x := i31 + 2 else ();
blanchet@58351
   334
            !x
blanchet@58351
   335
          end
blanchet@58351
   336
        else
blanchet@58351
   337
          if i31 + 1 < l andalso cmp (Array.sub (a, i31), Array.sub (a, i31 + 1)) = LESS
blanchet@58351
   338
          then i31 + 1 else if i31 < l then i31 else raise BOTTOM i
blanchet@58351
   339
      end
blanchet@58351
   340
blanchet@58351
   341
    fun trickledown l i e =
blanchet@58371
   342
      let val j = maxson l i in
blanchet@58351
   343
        if cmp (Array.sub (a, j), e) = GREATER then
blanchet@58371
   344
          (Array.update (a, i, Array.sub (a, j)); trickledown l j e)
blanchet@58371
   345
        else
blanchet@58371
   346
          Array.update (a, i, e)
blanchet@58351
   347
      end
blanchet@58351
   348
blanchet@58351
   349
    fun trickle l i e = trickledown l i e handle BOTTOM i => Array.update (a, i, e)
blanchet@58351
   350
blanchet@58351
   351
    fun bubbledown l i =
blanchet@58371
   352
      let val j = maxson l i in
blanchet@58371
   353
        Array.update (a, i, Array.sub (a, j));
blanchet@58351
   354
        bubbledown l j
blanchet@58351
   355
      end
blanchet@58351
   356
blanchet@58351
   357
    fun bubble l i = bubbledown l i handle BOTTOM i => i
blanchet@58351
   358
blanchet@58351
   359
    fun trickleup i e =
blanchet@58371
   360
      let val father = (i - 1) div 3 in
blanchet@58351
   361
        if cmp (Array.sub (a, father), e) = LESS then
blanchet@58371
   362
          (Array.update (a, i, Array.sub (a, father));
blanchet@58371
   363
           if father > 0 then trickleup father e else Array.update (a, 0, e))
blanchet@58371
   364
        else
blanchet@58371
   365
          Array.update (a, i, e)
blanchet@58351
   366
      end
blanchet@58351
   367
blanchet@58697
   368
    fun for i = if i < 0 then () else (trickle al i (Array.sub (a, i)); for (i - 1))
blanchet@58351
   369
blanchet@58351
   370
    fun for2 i =
blanchet@58697
   371
      if i < Integer.max 2 (al - bnd) then
blanchet@58371
   372
        ()
blanchet@58371
   373
      else
blanchet@58371
   374
        let val e = Array.sub (a, i) in
blanchet@58371
   375
          Array.update (a, i, Array.sub (a, 0));
blanchet@58371
   376
          trickleup (bubble i 0) e;
blanchet@58371
   377
          for2 (i - 1)
blanchet@58371
   378
        end
blanchet@58351
   379
  in
blanchet@58697
   380
    for (((al + 1) div 3) - 1);
blanchet@58697
   381
    for2 (al - 1);
blanchet@58697
   382
    if al > 1 then
blanchet@58371
   383
      let val e = Array.sub (a, 1) in
blanchet@58371
   384
        Array.update (a, 1, Array.sub (a, 0));
blanchet@58351
   385
        Array.update (a, 0, e)
blanchet@58351
   386
      end
blanchet@58371
   387
    else
blanchet@58371
   388
      ()
blanchet@58351
   389
  end
blanchet@58351
   390
blanchet@58695
   391
val number_of_nearest_neighbors = 10 (* FUDGE *)
blanchet@58695
   392
blanchet@58713
   393
fun select_visible_facts big_number recommends =
blanchet@58706
   394
  List.app (fn at =>
blanchet@58706
   395
    let val (j, ov) = Array.sub (recommends, at) in
blanchet@58713
   396
      Array.update (recommends, at, (j, big_number + ov))
blanchet@58706
   397
    end)
blanchet@58706
   398
blanchet@58695
   399
exception EXIT of unit
blanchet@58695
   400
blanchet@58718
   401
fun k_nearest_neighbors dffreq num_facts depss feat_facts max_suggs visible_facts goal_feats =
blanchet@58351
   402
  let
blanchet@58699
   403
    val ln_afreq = Math.ln (Real.fromInt num_facts)
blanchet@58716
   404
    fun tfidf feat = ln_afreq - Math.ln (Real.fromInt (Vector.sub (dffreq, feat)))
blanchet@58359
   405
blanchet@58444
   406
    val overlaps_sqr = Array.tabulate (num_facts, rpair 0.0)
blanchet@58359
   407
blanchet@58351
   408
    fun inc_overlap j v =
blanchet@58703
   409
      let val ov = snd (Array.sub (overlaps_sqr, j)) in
blanchet@58351
   410
        Array.update (overlaps_sqr, j, (j, v + ov))
blanchet@58359
   411
      end
blanchet@58359
   412
blanchet@58699
   413
    fun do_feat s =
blanchet@58351
   414
      let
blanchet@58699
   415
        val sw = tfidf s
blanchet@58699
   416
        val w2 = sw * sw
blanchet@58699
   417
        fun do_th j = if j < num_facts then inc_overlap j w2 else ()
blanchet@58351
   418
      in
blanchet@58716
   419
        List.app do_th (Array.sub (feat_facts, s))
blanchet@58359
   420
      end
blanchet@58359
   421
blanchet@58718
   422
    val _ = List.app do_feat goal_feats
blanchet@58697
   423
    val _ = heap (Real.compare o pairself snd) num_facts num_facts overlaps_sqr
blanchet@58695
   424
    val no_recommends = Unsynchronized.ref 0
blanchet@58698
   425
    val recommends = Array.tabulate (num_facts, rpair 0.0)
blanchet@58705
   426
    val age = Unsynchronized.ref 500000000.0
blanchet@58359
   427
blanchet@58351
   428
    fun inc_recommend j v =
blanchet@58695
   429
      let val ov = snd (Array.sub (recommends, j)) in
blanchet@58716
   430
        if ov <= 0.0 then
blanchet@58716
   431
          (no_recommends := !no_recommends + 1; Array.update (recommends, j, (j, !age + ov)))
blanchet@58716
   432
        else if ov < !age + 1000.0 then
blanchet@58716
   433
          Array.update (recommends, j, (j, v + ov))
blanchet@58716
   434
        else
blanchet@58716
   435
          ()
blanchet@58716
   436
      end
blanchet@58359
   437
blanchet@58695
   438
    val k = Unsynchronized.ref 0
blanchet@58695
   439
    fun do_k k =
blanchet@58698
   440
      if k >= num_facts then
blanchet@58695
   441
        raise EXIT ()
blanchet@58359
   442
      else
blanchet@58359
   443
        let
blanchet@58444
   444
          val (j, o2) = Array.sub (overlaps_sqr, num_facts - k - 1)
blanchet@58359
   445
          val o1 = Math.sqrt o2
blanchet@58359
   446
          val _ = inc_recommend j o1
blanchet@58716
   447
          val ds = Vector.sub (depss, j)
blanchet@58359
   448
          val l = Real.fromInt (length ds)
blanchet@58359
   449
        in
blanchet@58695
   450
          List.app (fn d => inc_recommend d (o1 / l)) ds
blanchet@58359
   451
        end
blanchet@58359
   452
blanchet@58695
   453
    fun while1 () =
blanchet@58703
   454
      if !k = number_of_nearest_neighbors then () else (do_k (!k); k := !k + 1; while1 ())
blanchet@58695
   455
      handle EXIT () => ()
blanchet@58695
   456
blanchet@58695
   457
    fun while2 () =
blanchet@58703
   458
      if !no_recommends >= max_suggs then ()
blanchet@58703
   459
      else (do_k (!k); k := !k + 1; age := !age - 10000.0; while2 ())
blanchet@58695
   460
      handle EXIT () => ()
blanchet@58695
   461
blanchet@58351
   462
    fun ret acc at =
blanchet@58705
   463
      if at = num_facts then acc else ret (Array.sub (recommends, at) :: acc) (at + 1)
blanchet@58351
   464
  in
blanchet@58706
   465
    while1 ();
blanchet@58706
   466
    while2 ();
blanchet@58713
   467
    select_visible_facts 1000000000.0 recommends visible_facts;
blanchet@58698
   468
    heap (Real.compare o pairself snd) max_suggs num_facts recommends;
blanchet@58698
   469
    ret [] (Integer.max 0 (num_facts - max_suggs))
blanchet@58351
   470
  end
blanchet@58351
   471
blanchet@58716
   472
fun wider_array_of_vector init vec =
blanchet@58716
   473
  let val ary = Array.array init in
blanchet@58716
   474
    Array.copyVec {src = vec, dst = ary, di = 0};
blanchet@58716
   475
    ary
blanchet@58716
   476
  end
blanchet@58716
   477
blanchet@58474
   478
val nb_def_prior_weight = 21 (* FUDGE *)
blanchet@58437
   479
blanchet@58716
   480
fun learn_facts (tfreq0, sfreq0, dffreq0) num_facts0 num_facts num_feats depss featss =
blanchet@58696
   481
  let
blanchet@58716
   482
    val tfreq = wider_array_of_vector (num_facts, 0) tfreq0
blanchet@58716
   483
    val sfreq = wider_array_of_vector (num_facts, Inttab.empty) sfreq0
blanchet@58716
   484
    val dffreq = wider_array_of_vector (num_feats, 0) dffreq0
blanchet@58716
   485
blanchet@58716
   486
    fun learn_one th feats deps =
blanchet@58696
   487
      let
blanchet@58697
   488
        fun add_th weight t =
blanchet@58697
   489
          let
blanchet@58697
   490
            val im = Array.sub (sfreq, t)
blanchet@58716
   491
            fun fold_fn s = Inttab.map_default (s, 0) (Integer.add weight)
blanchet@58697
   492
          in
blanchet@58710
   493
            map_array_at tfreq (Integer.add weight) t;
blanchet@58697
   494
            Array.update (sfreq, t, fold fold_fn feats im)
blanchet@58697
   495
          end
blanchet@58697
   496
blanchet@58710
   497
        val add_sym = map_array_at dffreq (Integer.add 1)
blanchet@58696
   498
      in
blanchet@58697
   499
        add_th nb_def_prior_weight th;
blanchet@58697
   500
        List.app (add_th 1) deps;
blanchet@58697
   501
        List.app add_sym feats
blanchet@58696
   502
      end
blanchet@58696
   503
blanchet@58444
   504
    fun for i =
blanchet@58725
   505
      if i = num_facts then ()
blanchet@58725
   506
      else (learn_one i (Vector.sub (featss, i)) (Vector.sub (depss, i)); for (i + 1))
blanchet@58444
   507
  in
blanchet@58725
   508
    for num_facts0;
blanchet@58716
   509
    (Array.vector tfreq, Array.vector sfreq, Array.vector dffreq)
blanchet@58444
   510
  end
blanchet@58371
   511
blanchet@58718
   512
fun naive_bayes (tfreq, sfreq, dffreq) num_facts max_suggs visible_facts goal_feats =
blanchet@58620
   513
  let
blanchet@58704
   514
    val tau = 0.05 (* FUDGE *)
blanchet@58704
   515
    val pos_weight = 10.0 (* FUDGE *)
blanchet@58620
   516
    val def_val = ~15.0 (* FUDGE *)
blanchet@58466
   517
blanchet@58708
   518
    val ln_afreq = Math.ln (Real.fromInt num_facts)
blanchet@58716
   519
    val idf = Vector.map (fn i => ln_afreq - Math.ln (Real.fromInt i)) dffreq
blanchet@58708
   520
blanchet@58701
   521
    fun tfidf feat = Vector.sub (idf, feat)
blanchet@58444
   522
blanchet@58444
   523
    fun log_posterior i =
blanchet@58371
   524
      let
blanchet@58716
   525
        val tfreq = Real.fromInt (Vector.sub (tfreq, i))
blanchet@58439
   526
blanchet@58699
   527
        fun fold_feats f (res, sfh) =
blanchet@58444
   528
          (case Inttab.lookup sfh f of
blanchet@58444
   529
            SOME sf =>
blanchet@58640
   530
            (res + tfidf f * Math.ln (pos_weight * Real.fromInt sf / tfreq),
blanchet@58444
   531
             Inttab.delete f sfh)
blanchet@58640
   532
          | NONE => (res + tfidf f * def_val, sfh))
blanchet@58439
   533
blanchet@58718
   534
        val (res, sfh) = fold fold_feats goal_feats (Math.ln tfreq, Vector.sub (sfreq, i))
blanchet@58439
   535
blanchet@58704
   536
        fun fold_sfh (f, sf) sow = sow + tfidf f * Math.ln (1.0 + (1.0 - Real.fromInt sf) / tfreq)
blanchet@58439
   537
blanchet@58444
   538
        val sum_of_weights = Inttab.fold fold_sfh sfh 0.0
blanchet@58371
   539
      in
blanchet@58620
   540
        res + tau * sum_of_weights
blanchet@58439
   541
      end
blanchet@58371
   542
blanchet@58698
   543
    val posterior = Array.tabulate (num_facts, (fn j => (j, log_posterior j)))
blanchet@58444
   544
blanchet@58706
   545
    fun ret at acc =
blanchet@58706
   546
      if at = num_facts then acc else ret (at + 1) (Array.sub (posterior, at) :: acc)
blanchet@58371
   547
  in
blanchet@58713
   548
    select_visible_facts 100000.0 posterior visible_facts;
blanchet@58698
   549
    heap (Real.compare o pairself snd) max_suggs num_facts posterior;
blanchet@58706
   550
    ret (Integer.max 0 (num_facts - max_suggs)) []
blanchet@58371
   551
  end
blanchet@58371
   552
blanchet@58467
   553
(* experimental *)
blanchet@58718
   554
fun naive_bayes_py ctxt overlord num_facts depss featss max_suggs goal_feats =
blanchet@58467
   555
  let
blanchet@58467
   556
    fun name_of_fact j = "f" ^ string_of_int j
blanchet@58467
   557
    fun fact_of_name s = the (Int.fromString (unprefix "f" s))
blanchet@58467
   558
    fun name_of_feature j = "F" ^ string_of_int j
blanchet@58467
   559
    fun parents_of j = if j = 0 then [] else [name_of_fact (j - 1)]
blanchet@58467
   560
blanchet@58715
   561
    val learns = map (fn j => (name_of_fact j, parents_of j,
blanchet@58715
   562
      map name_of_feature (Vector.sub (featss, j)),
blanchet@58715
   563
      map name_of_fact (Vector.sub (depss, j)))) (0 upto num_facts - 1)
blanchet@58698
   564
    val parents' = parents_of num_facts
blanchet@58718
   565
    val goal_feats' = map (rpair 1.0 o name_of_feature) goal_feats
blanchet@58467
   566
  in
blanchet@58467
   567
    MaSh_Py.unlearn ctxt overlord;
blanchet@58472
   568
    OS.Process.sleep (seconds 2.0); (* hack *)
blanchet@58718
   569
    MaSh_Py.query ctxt overlord max_suggs (learns, [], parents', goal_feats')
blanchet@58467
   570
    |> map (apfst fact_of_name)
blanchet@58467
   571
  end
blanchet@58467
   572
blanchet@58633
   573
(* experimental *)
blanchet@58718
   574
fun external_tool tool max_suggs learns goal_feats =
blanchet@58633
   575
  let
blanchet@58639
   576
    val ser = string_of_int (serial ()) (* poor person's attempt at thread-safety *)
blanchet@58639
   577
    val ocs = TextIO.openOut ("adv_syms" ^ ser)
blanchet@58639
   578
    val ocd = TextIO.openOut ("adv_deps" ^ ser)
blanchet@58639
   579
    val ocq = TextIO.openOut ("adv_seq" ^ ser)
blanchet@58639
   580
    val occ = TextIO.openOut ("adv_conj" ^ ser)
blanchet@58638
   581
blanchet@58633
   582
    fun os oc s = TextIO.output (oc, s)
blanchet@58638
   583
blanchet@58639
   584
    fun ol _ _ _ [] = ()
blanchet@58639
   585
      | ol _ f _ [e] = f e
blanchet@58633
   586
      | ol oc f sep (h :: t) = (f h; os oc sep; ol oc f sep t)
blanchet@58638
   587
blanchet@58636
   588
    fun do_learn (name, feats, deps) =
blanchet@58699
   589
      (os ocs name; os ocs ":"; ol ocs (os ocs o quote) ", " feats; os ocs "\n";
blanchet@58639
   590
       os ocd name; os ocd ":"; ol ocd (os ocd) " " deps; os ocd "\n"; os ocq name; os ocq "\n")
blanchet@58638
   591
blanchet@58633
   592
    fun forkexec no =
blanchet@58633
   593
      let
blanchet@58633
   594
        val cmd =
blanchet@58639
   595
          "~/misc/" ^ tool ^ " adv_syms" ^ ser ^ " adv_deps" ^ ser ^ " " ^ string_of_int no ^
blanchet@58639
   596
          " adv_seq" ^ ser ^ " < adv_conj" ^ ser
blanchet@58633
   597
      in
blanchet@58633
   598
        fst (Isabelle_System.bash_output cmd)
blanchet@58633
   599
        |> space_explode " "
blanchet@58636
   600
        |> filter_out (curry (op =) "")
blanchet@58633
   601
      end
blanchet@58633
   602
  in
blanchet@58718
   603
    (List.app do_learn learns; ol occ (os occ o quote) ", " goal_feats;
blanchet@58639
   604
     TextIO.closeOut ocs; TextIO.closeOut ocd; TextIO.closeOut ocq; TextIO.closeOut occ;
blanchet@58636
   605
     forkexec max_suggs)
blanchet@58633
   606
  end
blanchet@58633
   607
blanchet@58717
   608
val k_nearest_neighbors_ext =
blanchet@58718
   609
  external_tool ("newknn/knn" ^ " " ^ string_of_int number_of_nearest_neighbors)
blanchet@58718
   610
val naive_bayes_ext = external_tool "predict/nbayes"
blanchet@58639
   611
blanchet@58718
   612
fun query_external ctxt engine max_suggs learns goal_feats =
blanchet@58718
   613
  (trace_msg ctxt (fn () => "MaSh_SML query external " ^ encode_strs goal_feats);
blanchet@58718
   614
   (case engine of
blanchet@58718
   615
     MaSh_SML_kNN_Ext => k_nearest_neighbors_ext max_suggs learns goal_feats
blanchet@58718
   616
   | MaSh_SML_NB_Ext => naive_bayes_ext max_suggs learns goal_feats))
blanchet@58716
   617
blanchet@58720
   618
fun query_internal ctxt engine num_facts num_feats (fact_names, featss, depss)
blanchet@58720
   619
    (freqs as (_, _, dffreq)) visible_facts max_suggs goal_feats int_goal_feats =
blanchet@58718
   620
  (trace_msg ctxt (fn () => "MaSh_SML query internal " ^ encode_strs goal_feats ^ " from {" ^
blanchet@58720
   621
     elide_string 1000 (space_implode " " (Vector.foldr (op ::) [] fact_names)) ^ "}");
blanchet@58718
   622
   (case engine of
blanchet@58718
   623
     MaSh_SML_kNN =>
blanchet@58718
   624
     let
blanchet@58718
   625
       val feat_facts = Array.array (num_feats, [])
blanchet@58718
   626
       val _ =
blanchet@58718
   627
         Vector.foldl (fn (feats, fact) =>
blanchet@58718
   628
             (List.app (map_array_at feat_facts (cons fact)) feats; fact + 1))
blanchet@58718
   629
           0 featss
blanchet@58718
   630
     in
blanchet@58718
   631
       k_nearest_neighbors dffreq num_facts depss feat_facts max_suggs visible_facts int_goal_feats
blanchet@58718
   632
     end
blanchet@58718
   633
   | MaSh_SML_NB => naive_bayes freqs num_facts max_suggs visible_facts int_goal_feats)
blanchet@58720
   634
   |> map (curry Vector.sub fact_names o fst))
blanchet@58349
   635
blanchet@58349
   636
end;
blanchet@58349
   637
blanchet@58349
   638
blanchet@51326
   639
(*** Middle-level communication with MaSh ***)
blanchet@51326
   640
blanchet@58347
   641
datatype proof_kind = Isar_Proof | Automatic_Proof | Isar_Proof_wegen_Prover_Flop
blanchet@51326
   642
blanchet@51326
   643
fun str_of_proof_kind Isar_Proof = "i"
blanchet@51499
   644
  | str_of_proof_kind Automatic_Proof = "a"
blanchet@51499
   645
  | str_of_proof_kind Isar_Proof_wegen_Prover_Flop = "x"
blanchet@51326
   646
blanchet@58347
   647
fun proof_kind_of_str "a" = Automatic_Proof
blanchet@51499
   648
  | proof_kind_of_str "x" = Isar_Proof_wegen_Prover_Flop
blanchet@58347
   649
  | proof_kind_of_str _ (* "i" *) = Isar_Proof
blanchet@51326
   650
blanchet@58353
   651
fun add_edge_to name parent =
blanchet@58353
   652
  Graph.default_node (parent, (Isar_Proof, [], []))
blanchet@58353
   653
  #> Graph.add_edge (parent, name)
blanchet@58353
   654
blanchet@58721
   655
fun add_node kind name parents feats deps (access_G, (fact_xtab, feat_xtab), learns) =
blanchet@58713
   656
  ((Graph.new_node (name, (kind, feats, deps)) access_G
blanchet@58713
   657
    handle Graph.DUP _ => Graph.map_node name (K (kind, feats, deps)) access_G)
blanchet@58713
   658
   |> fold (add_edge_to name) parents,
blanchet@58725
   659
  (add_to_xtab name fact_xtab, fold maybe_add_to_xtab feats feat_xtab),
blanchet@58721
   660
  (name, feats, deps) :: learns)
blanchet@58353
   661
blanchet@51326
   662
fun try_graph ctxt when def f =
blanchet@51326
   663
  f ()
blanchet@58347
   664
  handle
blanchet@58347
   665
    Graph.CYCLES (cycle :: _) =>
blanchet@58347
   666
    (trace_msg ctxt (fn () => "Cycle involving " ^ commas cycle ^ " when " ^ when); def)
blanchet@58347
   667
  | Graph.DUP name =>
blanchet@58347
   668
    (trace_msg ctxt (fn () => "Duplicate fact " ^ quote name ^ " when " ^ when); def)
blanchet@58347
   669
  | Graph.UNDEF name =>
blanchet@58347
   670
    (trace_msg ctxt (fn () => "Unknown fact " ^ quote name ^ " when " ^ when); def)
blanchet@58347
   671
  | exn =>
blanchet@58347
   672
    if Exn.is_interrupt exn then
blanchet@58347
   673
      reraise exn
blanchet@58347
   674
    else
blanchet@58347
   675
      (trace_msg ctxt (fn () => "Internal error when " ^ when ^ ":\n" ^ Runtime.exn_message exn);
blanchet@58347
   676
       def)
blanchet@51326
   677
blanchet@51326
   678
fun graph_info G =
blanchet@51326
   679
  string_of_int (length (Graph.keys G)) ^ " node(s), " ^
blanchet@58348
   680
  string_of_int (fold (Integer.add o length o snd) (Graph.dest G) 0) ^ " edge(s), " ^
blanchet@51326
   681
  string_of_int (length (Graph.maximals G)) ^ " maximal"
blanchet@51326
   682
blanchet@54232
   683
type mash_state =
blanchet@58700
   684
  {access_G : (proof_kind * string list * string list) Graph.T,
blanchet@58716
   685
   xtabs : xtab * xtab,
blanchet@58720
   686
   ffds : string vector * int list vector * int list vector,
blanchet@58720
   687
   freqs : int vector * int Inttab.table vector * int vector,
blanchet@58707
   688
   dirty_facts : string list option}
blanchet@51326
   689
blanchet@58720
   690
val empty_xtabs = (empty_xtab, empty_xtab)
blanchet@58720
   691
val empty_ffds = (Vector.fromList [], Vector.fromList [], Vector.fromList [])
blanchet@58720
   692
val empty_freqs = (Vector.fromList [], Vector.fromList [], Vector.fromList [])
blanchet@58720
   693
blanchet@58713
   694
val empty_state =
blanchet@58713
   695
  {access_G = Graph.empty,
blanchet@58720
   696
   xtabs = empty_xtabs,
blanchet@58720
   697
   ffds = empty_ffds,
blanchet@58720
   698
   freqs = empty_freqs,
blanchet@58713
   699
   dirty_facts = SOME []} : mash_state
blanchet@58713
   700
blanchet@58722
   701
fun recompute_ffds_freqs_from_learns learns ((num_facts, fact_tab), (num_feats, feat_tab))
blanchet@58722
   702
    num_facts0 (fact_names0, featss0, depss0) freqs0 =
blanchet@58721
   703
  let
blanchet@58722
   704
    val fact_names = Vector.concat [fact_names0, Vector.fromList (map #1 learns)]
blanchet@58722
   705
    val featss = Vector.concat [featss0,
blanchet@58722
   706
      Vector.fromList (map (map_filter (Symtab.lookup feat_tab) o #2) learns)]
blanchet@58722
   707
    val depss = Vector.concat [depss0,
blanchet@58722
   708
      Vector.fromList (map (map_filter (Symtab.lookup fact_tab) o #3) learns)]
blanchet@58721
   709
  in
blanchet@58721
   710
    ((fact_names, featss, depss),
blanchet@58722
   711
     MaSh_SML.learn_facts freqs0 num_facts0 num_facts num_feats depss featss)
blanchet@58721
   712
  end
blanchet@58721
   713
blanchet@58720
   714
fun reorder_learns (num_facts, fact_tab) learns =
blanchet@58720
   715
  let val ary = Array.array (num_facts, ("", [], [])) in
blanchet@58720
   716
    List.app (fn learn as (fact, _, _) =>
blanchet@58720
   717
        Array.update (ary, the (Symtab.lookup fact_tab fact), learn))
blanchet@58720
   718
      learns;
blanchet@58720
   719
    Array.foldr (op ::) [] ary
blanchet@58720
   720
  end
blanchet@58720
   721
blanchet@58722
   722
fun recompute_ffds_freqs_from_access_G access_G (xtabs as (fact_xtab, _)) =
blanchet@58720
   723
  let
blanchet@58720
   724
    val learns =
blanchet@58720
   725
      Graph.schedule (fn _ => fn (fact, (_, feats, deps)) => (fact, feats, deps)) access_G
blanchet@58720
   726
      |> reorder_learns fact_xtab
blanchet@58720
   727
  in
blanchet@58722
   728
    recompute_ffds_freqs_from_learns learns xtabs 0 empty_ffds empty_freqs
blanchet@58720
   729
  end
blanchet@51326
   730
blanchet@51326
   731
local
blanchet@51326
   732
blanchet@58700
   733
val version = "*** MaSh version 20140625 ***"
blanchet@51372
   734
blanchet@54232
   735
exception FILE_VERSION_TOO_NEW of unit
blanchet@51326
   736
blanchet@51326
   737
fun extract_node line =
blanchet@56628
   738
  (case space_explode ":" line of
blanchet@58347
   739
    [head, tail] =>
blanchet@58347
   740
    (case (space_explode " " head, map (unprefix " ") (space_explode ";" tail)) of
blanchet@58347
   741
      ([kind, name], [parents, feats, deps]) =>
blanchet@58700
   742
      SOME (proof_kind_of_str kind, decode_str name, decode_strs parents, decode_strs feats,
blanchet@58352
   743
        decode_strs deps)
blanchet@56628
   744
    | _ => NONE)
blanchet@56628
   745
  | _ => NONE)
blanchet@51326
   746
blanchet@58418
   747
fun load_state ctxt overlord (time_state as (memory_time, _)) =
blanchet@58418
   748
  let val path = mash_state_file () in
blanchet@58728
   749
    (case try OS.FileSys.modTime (Path.implode path) of
blanchet@58418
   750
      NONE => time_state
blanchet@58418
   751
    | SOME disk_time =>
blanchet@58418
   752
      if Time.>= (memory_time, disk_time) then
blanchet@58418
   753
        time_state
blanchet@58418
   754
      else
blanchet@58418
   755
        (disk_time,
blanchet@58418
   756
         (case try File.read_lines path of
blanchet@58418
   757
           SOME (version' :: node_lines) =>
blanchet@58418
   758
           let
blanchet@58418
   759
             fun extract_line_and_add_node line =
blanchet@58418
   760
               (case extract_node line of
blanchet@58418
   761
                 NONE => I (* should not happen *)
blanchet@58418
   762
               | SOME (kind, name, parents, feats, deps) => add_node kind name parents feats deps)
blanchet@58353
   763
blanchet@58721
   764
             val empty_G_etc = (Graph.empty, empty_xtabs, [])
blanchet@58721
   765
blanchet@58721
   766
             val (access_G, xtabs, rev_learns) =
blanchet@58418
   767
               (case string_ord (version', version) of
blanchet@58418
   768
                 EQUAL =>
blanchet@58721
   769
                 try_graph ctxt "loading state" empty_G_etc
blanchet@58721
   770
                   (fn () => fold extract_line_and_add_node node_lines empty_G_etc)
blanchet@58418
   771
               | LESS =>
blanchet@58418
   772
                 (* cannot parse old file *)
blanchet@58418
   773
                 (if the_mash_engine () = MaSh_Py then MaSh_Py.unlearn ctxt overlord
blanchet@58418
   774
                  else wipe_out_mash_state_dir ();
blanchet@58721
   775
                  empty_G_etc)
blanchet@58418
   776
               | GREATER => raise FILE_VERSION_TOO_NEW ())
blanchet@58720
   777
blanchet@58722
   778
             val (ffds, freqs) =
blanchet@58722
   779
               recompute_ffds_freqs_from_learns (rev rev_learns) xtabs 0 empty_ffds empty_freqs
blanchet@58418
   780
           in
blanchet@58418
   781
             trace_msg ctxt (fn () => "Loaded fact graph (" ^ graph_info access_G ^ ")");
blanchet@58720
   782
             {access_G = access_G, xtabs = xtabs, ffds = ffds, freqs = freqs, dirty_facts = SOME []}
blanchet@58418
   783
           end
blanchet@58418
   784
         | _ => empty_state)))
blanchet@58418
   785
  end
blanchet@51326
   786
blanchet@58347
   787
fun str_of_entry (kind, name, parents, feats, deps) =
blanchet@58347
   788
  str_of_proof_kind kind ^ " " ^ encode_str name ^ ": " ^ encode_strs parents ^ "; " ^
blanchet@58700
   789
  encode_strs feats ^ "; " ^ encode_strs deps ^ "\n"
blanchet@58347
   790
blanchet@58707
   791
fun save_state _ (time_state as (_, {dirty_facts = SOME [], ...})) = time_state
blanchet@58720
   792
  | save_state ctxt (memory_time, {access_G, xtabs, ffds, freqs, dirty_facts}) =
blanchet@51326
   793
    let
blanchet@58347
   794
      fun append_entry (name, ((kind, feats, deps), (parents, _))) =
blanchet@58347
   795
        cons (kind, name, Graph.Keys.dest parents, feats, deps)
blanchet@58347
   796
blanchet@58418
   797
      val path = mash_state_file ()
blanchet@58707
   798
      val dirty_facts' =
blanchet@58418
   799
        (case try OS.FileSys.modTime (Path.implode path) of
blanchet@58418
   800
          NONE => NONE
blanchet@58707
   801
        | SOME disk_time => if Time.< (disk_time, memory_time) then dirty_facts else NONE)
blanchet@51326
   802
      val (banner, entries) =
blanchet@58707
   803
        (case dirty_facts' of
blanchet@56628
   804
          SOME names => (NONE, fold (append_entry o Graph.get_entry access_G) names [])
blanchet@56628
   805
        | NONE => (SOME (version ^ "\n"), Graph.fold append_entry access_G []))
blanchet@51326
   806
    in
blanchet@58418
   807
      write_file banner (entries, str_of_entry) path;
blanchet@51326
   808
      trace_msg ctxt (fn () =>
blanchet@58347
   809
        "Saved fact graph (" ^ graph_info access_G ^
blanchet@58707
   810
        (case dirty_facts of
blanchet@58707
   811
          SOME dirty_facts => "; " ^ string_of_int (length dirty_facts) ^ " dirty fact(s)"
blanchet@58347
   812
        | _ => "") ^  ")");
blanchet@58720
   813
      (Time.now (),
blanchet@58720
   814
       {access_G = access_G, xtabs = xtabs, ffds = ffds, freqs = freqs, dirty_facts = SOME []})
blanchet@51326
   815
    end
blanchet@51326
   816
blanchet@58707
   817
val global_state = Synchronized.var "Sledgehammer_MaSh.global_state" (Time.zeroTime, empty_state)
blanchet@51326
   818
blanchet@51326
   819
in
blanchet@51326
   820
blanchet@54285
   821
fun map_state ctxt overlord f =
blanchet@58418
   822
  Synchronized.change global_state (load_state ctxt overlord ##> f #> save_state ctxt)
blanchet@54232
   823
  handle FILE_VERSION_TOO_NEW () => ()
blanchet@51326
   824
blanchet@54285
   825
fun peek_state ctxt overlord f =
blanchet@58347
   826
  Synchronized.change_result global_state (perhaps (try (load_state ctxt overlord)) #> `snd #>> f)
blanchet@51326
   827
blanchet@54285
   828
fun clear_state ctxt overlord =
blanchet@58359
   829
  (* "MaSh_Py.unlearn" also removes the state file *)
blanchet@58349
   830
  Synchronized.change global_state (fn _ =>
blanchet@58371
   831
    (if the_mash_engine () = MaSh_Py then MaSh_Py.unlearn ctxt overlord
blanchet@58371
   832
     else wipe_out_mash_state_dir ();
blanchet@58418
   833
     (Time.zeroTime, empty_state)))
blanchet@51326
   834
blanchet@51326
   835
end
blanchet@51326
   836
blanchet@54695
   837
fun mash_unlearn ctxt ({overlord, ...} : params) =
blanchet@54696
   838
  (clear_state ctxt overlord; Output.urgent_message "Reset MaSh.")
blanchet@51585
   839
blanchet@51326
   840
blanchet@51326
   841
(*** Isabelle helpers ***)
blanchet@51326
   842
blanchet@51737
   843
val local_prefix = "local" ^ Long_Name.separator
blanchet@49393
   844
blanchet@51737
   845
fun elided_backquote_thm threshold th =
blanchet@58348
   846
  elide_string threshold (backquote_thm (Proof_Context.init_global (Thm.theory_of_thm th)) th)
blanchet@49393
   847
blanchet@52318
   848
val thy_name_of_thm = Context.theory_name o Thm.theory_of_thm
blanchet@52318
   849
blanchet@51639
   850
fun nickname_of_thm th =
blanchet@49409
   851
  if Thm.has_name_hint th then
blanchet@49409
   852
    let val hint = Thm.get_name_hint th in
blanchet@51737
   853
      (* There must be a better way to detect local facts. *)
blanchet@56628
   854
      (case try (unprefix local_prefix) hint of
blanchet@49409
   855
        SOME suf =>
blanchet@56628
   856
        thy_name_of_thm th ^ Long_Name.separator ^ suf ^ Long_Name.separator ^
blanchet@56628
   857
        elided_backquote_thm 50 th
blanchet@56628
   858
      | NONE => hint)
blanchet@49409
   859
    end
blanchet@49409
   860
  else
blanchet@51737
   861
    elided_backquote_thm 200 th
blanchet@49393
   862
blanchet@52271
   863
fun find_suggested_facts ctxt facts =
blanchet@49345
   864
  let
blanchet@52271
   865
    fun add (fact as (_, th)) = Symtab.default (nickname_of_thm th, fact)
blanchet@52271
   866
    val tab = fold add facts Symtab.empty
blanchet@52271
   867
    fun lookup nick =
blanchet@52271
   868
      Symtab.lookup tab nick
blanchet@58348
   869
      |> tap (fn NONE => trace_msg ctxt (fn () => "Cannot find " ^ quote nick) | _ => ())
blanchet@52271
   870
  in map_filter lookup end
blanchet@49326
   871
blanchet@51398
   872
fun scaled_avg [] = 0
blanchet@58348
   873
  | scaled_avg xs = Real.ceil (100000000.0 * fold (curry (op +)) xs 0.0) div length xs
blanchet@49343
   874
blanchet@51398
   875
fun avg [] = 0.0
blanchet@51398
   876
  | avg xs = fold (curry (op +)) xs 0.0 / Real.fromInt (length xs)
blanchet@49328
   877
blanchet@51398
   878
fun normalize_scores _ [] = []
blanchet@51398
   879
  | normalize_scores max_facts xs =
blanchet@58348
   880
    map (apsnd (curry Real.* (1.0 / avg (map snd (take max_facts xs))))) xs
blanchet@51398
   881
blanchet@58438
   882
fun mesh_facts fact_eq max_facts [(_, (sels, unks))] =
blanchet@58438
   883
    distinct fact_eq (map fst (take max_facts sels) @ take (max_facts - length sels) unks)
blanchet@51876
   884
  | mesh_facts fact_eq max_facts mess =
blanchet@49329
   885
    let
blanchet@52211
   886
      val mess = mess |> map (apsnd (apfst (normalize_scores max_facts)))
blanchet@58348
   887
blanchet@52211
   888
      fun score_in fact (global_weight, (sels, unks)) =
blanchet@58348
   889
        let val score_at = try (nth sels) #> Option.map (fn (_, score) => global_weight * score) in
blanchet@56628
   890
          (case find_index (curry fact_eq fact o fst) sels of
blanchet@55541
   891
            ~1 => if member fact_eq unks fact then NONE else SOME 0.0
blanchet@56628
   892
          | rank => score_at rank)
blanchet@51398
   893
        end
blanchet@58348
   894
blanchet@51398
   895
      fun weight_of fact = mess |> map_filter (score_in fact) |> scaled_avg
blanchet@49329
   896
    in
blanchet@58348
   897
      fold (union fact_eq o map fst o take max_facts o fst o snd) mess []
blanchet@58348
   898
      |> map (`weight_of) |> sort (int_ord o swap o pairself fst)
blanchet@58348
   899
      |> map snd |> take max_facts
blanchet@49329
   900
    end
blanchet@49327
   901
blanchet@56035
   902
val default_weight = 1.0
blanchet@58397
   903
fun free_feature_of s = ("f" ^ s, 40.0 (* FUDGE *))
blanchet@58397
   904
fun thy_feature_of s = ("y" ^ s, 8.0 (* FUDGE *))
blanchet@58397
   905
fun type_feature_of s = ("t" ^ s, 4.0 (* FUDGE *))
blanchet@58397
   906
fun class_feature_of s = ("s" ^ s, 1.0 (* FUDGE *))
blanchet@58397
   907
val local_feature = ("local", 16.0 (* FUDGE *))
blanchet@49266
   908
blanchet@52272
   909
fun crude_theory_ord p =
blanchet@51737
   910
  if Theory.subthy p then
blanchet@51737
   911
    if Theory.eq_thy p then EQUAL else LESS
blanchet@49339
   912
  else if Theory.subthy (swap p) then
blanchet@49339
   913
    GREATER
blanchet@56628
   914
  else
blanchet@56628
   915
    (case int_ord (pairself (length o Theory.ancestors_of) p) of
blanchet@56628
   916
      EQUAL => string_ord (pairself Context.theory_name p)
blanchet@56628
   917
    | order => order)
blanchet@49339
   918
blanchet@52272
   919
fun crude_thm_ord p =
blanchet@56628
   920
  (case crude_theory_ord (pairself theory_of_thm p) of
blanchet@51374
   921
    EQUAL =>
blanchet@58381
   922
    (* The hack below is necessary because of odd dependencies that are not reflected in the theory
blanchet@58381
   923
       comparison. *)
blanchet@51639
   924
    let val q = pairself nickname_of_thm p in
blanchet@51639
   925
      (* Hack to put "xxx_def" before "xxxI" and "xxxE" *)
blanchet@56628
   926
      (case bool_ord (pairself (String.isSuffix "_def") (swap q)) of
blanchet@51639
   927
        EQUAL => string_ord q
blanchet@56628
   928
      | ord => ord)
blanchet@51639
   929
    end
blanchet@56628
   930
  | ord => ord)
blanchet@49339
   931
blanchet@52273
   932
val thm_less_eq = Theory.subthy o pairself theory_of_thm
blanchet@52273
   933
fun thm_less p = thm_less_eq p andalso not (thm_less_eq (swap p))
blanchet@52273
   934
blanchet@49407
   935
val freezeT = Type.legacy_freeze_type
blanchet@49407
   936
blanchet@49407
   937
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49407
   938
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49407
   939
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49407
   940
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49407
   941
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49407
   942
  | freeze t = t
blanchet@49407
   943
blanchet@49407
   944
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
blanchet@49407
   945
blanchet@55593
   946
fun run_prover_for_mash ctxt params prover goal_name facts goal =
blanchet@49407
   947
  let
blanchet@49407
   948
    val problem =
blanchet@55593
   949
      {comment = "Goal: " ^ goal_name, state = Proof.init ctxt, goal = goal, subgoal = 1,
blanchet@55593
   950
       subgoal_count = 1, factss = [("", facts)]}
blanchet@49407
   951
  in
blanchet@55876
   952
    get_minimizing_prover ctxt MaSh (K ()) prover params (K (K (K ""))) problem
blanchet@49407
   953
  end
blanchet@49407
   954
blanchet@49341
   955
val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49341
   956
blanchet@54223
   957
val pat_tvar_prefix = "_"
blanchet@54223
   958
val pat_var_prefix = "_"
blanchet@54220
   959
blanchet@54226
   960
(* try "Long_Name.base_name" for shorter names *)
blanchet@58397
   961
fun massage_long_name s = s
blanchet@54223
   962
blanchet@58348
   963
val crude_str_of_sort = space_implode ":" o map massage_long_name o subtract (op =) @{sort type}
blanchet@54223
   964
blanchet@54223
   965
fun crude_str_of_typ (Type (s, [])) = massage_long_name s
blanchet@58348
   966
  | crude_str_of_typ (Type (s, Ts)) = massage_long_name s ^ implode (map crude_str_of_typ Ts)
blanchet@54220
   967
  | crude_str_of_typ (TFree (_, S)) = crude_str_of_sort S
blanchet@54220
   968
  | crude_str_of_typ (TVar (_, S)) = crude_str_of_sort S
blanchet@54220
   969
blanchet@54265
   970
fun maybe_singleton_str _ "" = []
blanchet@54265
   971
  | maybe_singleton_str pref s = [pref ^ s]
blanchet@54265
   972
blanchet@54285
   973
val max_pat_breadth = 10 (* FUDGE *)
blanchet@51600
   974
blanchet@58397
   975
fun term_features_of ctxt thy_name num_facts const_tab term_max_depth type_max_depth ts =
blanchet@49266
   976
  let
blanchet@51407
   977
    val thy = Proof_Context.theory_of ctxt
blanchet@54219
   978
blanchet@51408
   979
    val fixes = map snd (Variable.dest_fixes ctxt)
blanchet@51407
   980
    val classes = Sign.classes_of thy
blanchet@54219
   981
blanchet@49319
   982
    fun add_classes @{sort type} = I
blanchet@51407
   983
      | add_classes S =
blanchet@51407
   984
        fold (`(Sorts.super_classes classes)
blanchet@58348
   985
          #> swap #> op ::
blanchet@58348
   986
          #> subtract (op =) @{sort type} #> map massage_long_name
blanchet@58348
   987
          #> map class_feature_of
blanchet@58348
   988
          #> union (eq_fst (op =))) S
blanchet@54219
   989
blanchet@54219
   990
    fun pattify_type 0 _ = []
blanchet@54219
   991
      | pattify_type _ (Type (s, [])) =
blanchet@54223
   992
        if member (op =) bad_types s then [] else [massage_long_name s]
blanchet@54219
   993
      | pattify_type depth (Type (s, U :: Ts)) =
blanchet@54219
   994
        let
blanchet@54219
   995
          val T = Type (s, Ts)
blanchet@58397
   996
          val ps = take max_pat_breadth (pattify_type depth T)
blanchet@58397
   997
          val qs = take max_pat_breadth ("" :: pattify_type (depth - 1) U)
blanchet@58348
   998
        in
blanchet@58348
   999
          map_product (fn p => fn "" => p | q => p ^ "(" ^ q ^ ")") ps qs
blanchet@58348
  1000
        end
blanchet@54265
  1001
      | pattify_type _ (TFree (_, S)) =
blanchet@54265
  1002
        maybe_singleton_str pat_tvar_prefix (crude_str_of_sort S)
blanchet@54265
  1003
      | pattify_type _ (TVar (_, S)) =
blanchet@54265
  1004
        maybe_singleton_str pat_tvar_prefix (crude_str_of_sort S)
blanchet@58348
  1005
blanchet@54219
  1006
    fun add_type_pat depth T =
blanchet@54296
  1007
      union (eq_fst (op =)) (map type_feature_of (pattify_type depth T))
blanchet@58348
  1008
blanchet@54219
  1009
    fun add_type_pats 0 _ = I
blanchet@54219
  1010
      | add_type_pats depth t =
blanchet@54219
  1011
        add_type_pat depth t #> add_type_pats (depth - 1) t
blanchet@58348
  1012
blanchet@54220
  1013
    fun add_type T =
blanchet@54220
  1014
      add_type_pats type_max_depth T
blanchet@54293
  1015
      #> fold_atyps_sorts (add_classes o snd) T
blanchet@58348
  1016
blanchet@54221
  1017
    fun add_subtypes (T as Type (_, Ts)) = add_type T #> fold add_subtypes Ts
blanchet@54221
  1018
      | add_subtypes T = add_type T
blanchet@54219
  1019
blanchet@58441
  1020
    val base_weight_of_const = 16.0 (* FUDGE *)
blanchet@58441
  1021
    val weight_of_const =
blanchet@58441
  1022
      (if num_facts = 0 orelse Symtab.is_empty const_tab then
blanchet@58441
  1023
         K base_weight_of_const
blanchet@54267
  1024
       else
blanchet@58441
  1025
         fn s =>
blanchet@54267
  1026
         let val count = Symtab.lookup const_tab s |> the_default 1 in
blanchet@58441
  1027
           base_weight_of_const + Real.fromInt num_facts / Real.fromInt count
blanchet@54267
  1028
         end)
blanchet@58348
  1029
blanchet@58397
  1030
    fun pattify_term _ 0 _ = []
blanchet@58397
  1031
      | pattify_term _ _ (Const (s, _)) =
blanchet@58397
  1032
        if is_widely_irrelevant_const s then [] else [(massage_long_name s, weight_of_const s)]
blanchet@55541
  1033
      | pattify_term _ _ (Free (s, T)) =
blanchet@54265
  1034
        maybe_singleton_str pat_var_prefix (crude_str_of_typ T)
blanchet@58397
  1035
        |> map (rpair 1.0)
blanchet@54227
  1036
        |> (if member (op =) fixes s then
blanchet@58397
  1037
              cons (free_feature_of (massage_long_name
blanchet@58397
  1038
                  (thy_name ^ Long_Name.separator ^ s)))
blanchet@54227
  1039
            else
blanchet@54227
  1040
              I)
blanchet@55541
  1041
      | pattify_term _ _ (Var (_, T)) =
blanchet@58397
  1042
        maybe_singleton_str pat_var_prefix (crude_str_of_typ T) |> map (rpair default_weight)
blanchet@55541
  1043
      | pattify_term Ts _ (Bound j) =
blanchet@56037
  1044
        maybe_singleton_str pat_var_prefix (crude_str_of_typ (nth Ts j))
blanchet@58397
  1045
        |> map (rpair default_weight)
blanchet@55541
  1046
      | pattify_term Ts depth (t $ u) =
blanchet@51354
  1047
        let
blanchet@58397
  1048
          val ps = take max_pat_breadth (pattify_term Ts depth t)
blanchet@58397
  1049
          val qs = take max_pat_breadth (("", default_weight) :: pattify_term Ts (depth - 1) u)
blanchet@54267
  1050
        in
blanchet@58397
  1051
          map_product (fn ppw as (p, pw) =>
blanchet@58397
  1052
            fn ("", _) => ppw
blanchet@58397
  1053
             | (q, qw) => (p ^ "(" ^ q ^ ")", pw + qw)) ps qs
blanchet@54267
  1054
        end
blanchet@55541
  1055
      | pattify_term _ _ _ = []
blanchet@58348
  1056
blanchet@55541
  1057
    fun add_term_pat Ts = union (eq_fst (op =)) oo pattify_term Ts
blanchet@58348
  1058
blanchet@54267
  1059
    fun add_term_pats _ 0 _ = I
blanchet@58397
  1060
      | add_term_pats Ts depth t = add_term_pat Ts depth t #> add_term_pats Ts (depth - 1) t
blanchet@58348
  1061
blanchet@54267
  1062
    fun add_term Ts = add_term_pats Ts term_max_depth
blanchet@58348
  1063
blanchet@54222
  1064
    fun add_subterms Ts t =
blanchet@56628
  1065
      (case strip_comb t of
blanchet@55541
  1066
        (Const (s, T), args) =>
blanchet@55541
  1067
        (not (is_widely_irrelevant_const s) ? add_term Ts t)
blanchet@55541
  1068
        #> add_subtypes T
blanchet@55541
  1069
        #> fold (add_subterms Ts) args
blanchet@51872
  1070
      | (head, args) =>
blanchet@49266
  1071
        (case head of
blanchet@54267
  1072
           Free (_, T) => add_term Ts t #> add_subtypes T
blanchet@54221
  1073
         | Var (_, T) => add_subtypes T
blanchet@54222
  1074
         | Abs (_, T, body) => add_subtypes T #> add_subterms (T :: Ts) body
blanchet@49266
  1075
         | _ => I)
blanchet@56628
  1076
        #> fold (add_subterms Ts) args)
blanchet@58348
  1077
  in
blanchet@58348
  1078
    fold (add_subterms []) ts []
blanchet@58348
  1079
  end
blanchet@49266
  1080
blanchet@54222
  1081
val term_max_depth = 2
blanchet@54292
  1082
val type_max_depth = 1
blanchet@49266
  1083
blanchet@49266
  1084
(* TODO: Generate type classes for types? *)
blanchet@58397
  1085
fun features_of ctxt thy num_facts const_tab (scope, _) ts =
blanchet@51408
  1086
  let val thy_name = Context.theory_name thy in
blanchet@51408
  1087
    thy_feature_of thy_name ::
blanchet@58397
  1088
    term_features_of ctxt thy_name num_facts const_tab term_max_depth type_max_depth ts
blanchet@51408
  1089
    |> scope <> Global ? cons local_feature
blanchet@51408
  1090
  end
blanchet@49266
  1091
blanchet@58348
  1092
(* Too many dependencies is a sign that a decision procedure is at work. There is not much to learn
blanchet@58348
  1093
   from such proofs. *)
blanchet@51449
  1094
val max_dependencies = 20
blanchet@51499
  1095
blanchet@55577
  1096
val prover_default_max_facts = 25
blanchet@49266
  1097
blanchet@49453
  1098
(* "type_definition_xxx" facts are characterized by their use of "CollectI". *)
blanchet@51770
  1099
val typedef_dep = nickname_of_thm @{thm CollectI}
blanchet@58348
  1100
(* Mysterious parts of the class machinery create lots of proofs that refer exclusively to
blanchet@58348
  1101
   "someI_ex" (and to some internal constructions). *)
blanchet@51770
  1102
val class_some_dep = nickname_of_thm @{thm someI_ex}
blanchet@49453
  1103
blanchet@51843
  1104
val fundef_ths =
blanchet@58348
  1105
  @{thms fundef_ex1_existence fundef_ex1_uniqueness fundef_ex1_iff fundef_default_value}
blanchet@51843
  1106
  |> map nickname_of_thm
blanchet@51843
  1107
blanchet@49453
  1108
(* "Rep_xxx_inject", "Abs_xxx_inverse", etc., are derived using these facts. *)
blanchet@49453
  1109
val typedef_ths =
blanchet@58348
  1110
  @{thms type_definition.Abs_inverse type_definition.Rep_inverse type_definition.Rep
blanchet@58348
  1111
      type_definition.Rep_inject type_definition.Abs_inject type_definition.Rep_cases
blanchet@58348
  1112
      type_definition.Abs_cases type_definition.Rep_induct type_definition.Abs_induct
blanchet@58348
  1113
      type_definition.Rep_range type_definition.Abs_image}
blanchet@51639
  1114
  |> map nickname_of_thm
blanchet@49453
  1115
blanchet@49456
  1116
fun is_size_def [dep] th =
blanchet@56984
  1117
    (case first_field ".rec" dep of
blanchet@58348
  1118
      SOME (pref, _) =>
blanchet@58348
  1119
      (case first_field ".size" (nickname_of_thm th) of
blanchet@58348
  1120
        SOME (pref', _) => pref = pref'
blanchet@58348
  1121
      | NONE => false)
blanchet@58348
  1122
    | NONE => false)
blanchet@49456
  1123
  | is_size_def _ _ = false
blanchet@49456
  1124
blanchet@52314
  1125
fun trim_dependencies deps =
blanchet@51770
  1126
  if length deps > max_dependencies then NONE else SOME deps
blanchet@49266
  1127
blanchet@51770
  1128
fun isar_dependencies_of name_tabs th =
blanchet@58648
  1129
  thms_in_proof max_dependencies (SOME name_tabs) th
blanchet@58648
  1130
  |> Option.map (fn deps =>
blanchet@58348
  1131
    if deps = [typedef_dep] orelse deps = [class_some_dep] orelse
blanchet@58348
  1132
       exists (member (op =) fundef_ths) deps orelse exists (member (op =) typedef_ths) deps orelse
blanchet@51843
  1133
       is_size_def deps th then
blanchet@51770
  1134
      []
blanchet@51770
  1135
    else
blanchet@58648
  1136
      deps)
blanchet@49419
  1137
blanchet@58348
  1138
fun prover_dependencies_of ctxt (params as {verbose, max_facts, ...}) prover auto_level facts
blanchet@58348
  1139
    name_tabs th =
blanchet@56628
  1140
  (case isar_dependencies_of name_tabs th of
blanchet@58648
  1141
    SOME [] => (false, [])
blanchet@58648
  1142
  | isar_deps0 =>
blanchet@49407
  1143
    let
blanchet@58648
  1144
      val isar_deps = these isar_deps0
blanchet@49407
  1145
      val thy = Proof_Context.theory_of ctxt
blanchet@49407
  1146
      val goal = goal_of_thm thy th
blanchet@55593
  1147
      val name = nickname_of_thm th
blanchet@53333
  1148
      val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal 1 ctxt
blanchet@52273
  1149
      val facts = facts |> filter (fn (_, th') => thm_less (th', th))
blanchet@58348
  1150
blanchet@52186
  1151
      fun nickify ((_, stature), th) = ((nickname_of_thm th, stature), th)
blanchet@58348
  1152
blanchet@51639
  1153
      fun is_dep dep (_, th) = nickname_of_thm th = dep
blanchet@58348
  1154
blanchet@49407
  1155
      fun add_isar_dep facts dep accum =
blanchet@49407
  1156
        if exists (is_dep dep) accum then
blanchet@49407
  1157
          accum
blanchet@56628
  1158
        else
blanchet@56628
  1159
          (case find_first (is_dep dep) facts of
blanchet@56628
  1160
            SOME ((_, status), th) => accum @ [(("", status), th)]
blanchet@58347
  1161
          | NONE => accum (* should not happen *))
blanchet@58348
  1162
blanchet@55575
  1163
      val mepo_facts =
blanchet@51499
  1164
        facts
blanchet@55547
  1165
        |> mepo_suggested_facts ctxt params (max_facts |> the_default prover_default_max_facts) NONE
blanchet@55547
  1166
             hyp_ts concl_t
blanchet@55575
  1167
      val facts =
blanchet@55575
  1168
        mepo_facts
blanchet@51769
  1169
        |> fold (add_isar_dep facts) isar_deps
blanchet@51639
  1170
        |> map nickify
blanchet@55575
  1171
      val num_isar_deps = length isar_deps
blanchet@49407
  1172
    in
blanchet@49419
  1173
      if verbose andalso auto_level = 0 then
blanchet@58359
  1174
        Output.urgent_message ("MaSh: " ^ quote prover ^ " on " ^ quote name ^ " with " ^
blanchet@58359
  1175
          string_of_int num_isar_deps ^ " + " ^ string_of_int (length facts - num_isar_deps) ^
blanchet@58359
  1176
          " facts.")
blanchet@49407
  1177
      else
blanchet@49407
  1178
        ();
blanchet@56628
  1179
      (case run_prover_for_mash ctxt params prover name facts goal of
blanchet@49407
  1180
        {outcome = NONE, used_facts, ...} =>
blanchet@49419
  1181
        (if verbose andalso auto_level = 0 then
blanchet@49407
  1182
           let val num_facts = length used_facts in
blanchet@58359
  1183
             Output.urgent_message ("Found proof with " ^ string_of_int num_facts ^ " fact" ^
blanchet@58359
  1184
               plural_s num_facts ^ ".")
blanchet@49407
  1185
           end
blanchet@49407
  1186
         else
blanchet@49407
  1187
           ();
blanchet@51769
  1188
         (true, map fst used_facts))
blanchet@56628
  1189
      | _ => (false, isar_deps))
blanchet@56628
  1190
    end)
blanchet@49266
  1191
blanchet@49266
  1192
blanchet@49266
  1193
(*** High-level communication with MaSh ***)
blanchet@49266
  1194
blanchet@52319
  1195
(* In the following functions, chunks are risers w.r.t. "thm_less_eq". *)
blanchet@52319
  1196
blanchet@52318
  1197
fun chunks_and_parents_for chunks th =
blanchet@52318
  1198
  let
blanchet@52318
  1199
    fun insert_parent new parents =
blanchet@52318
  1200
      let val parents = parents |> filter_out (fn p => thm_less_eq (p, new)) in
blanchet@58438
  1201
        parents |> forall (fn p => not (thm_less_eq (new, p))) parents ? cons new
blanchet@52318
  1202
      end
blanchet@58348
  1203
blanchet@52318
  1204
    fun rechunk seen (rest as th' :: ths) =
blanchet@52318
  1205
      if thm_less_eq (th', th) then (rev seen, rest)
blanchet@52318
  1206
      else rechunk (th' :: seen) ths
blanchet@58348
  1207
blanchet@52318
  1208
    fun do_chunk [] accum = accum
blanchet@52318
  1209
      | do_chunk (chunk as hd_chunk :: _) (chunks, parents) =
blanchet@52318
  1210
        if thm_less_eq (hd_chunk, th) then
blanchet@52318
  1211
          (chunk :: chunks, insert_parent hd_chunk parents)
blanchet@52318
  1212
        else if thm_less_eq (List.last chunk, th) then
blanchet@52318
  1213
          let val (front, back as hd_back :: _) = rechunk [] chunk in
blanchet@52318
  1214
            (front :: back :: chunks, insert_parent hd_back parents)
blanchet@52318
  1215
          end
blanchet@52318
  1216
        else
blanchet@52318
  1217
          (chunk :: chunks, parents)
blanchet@52318
  1218
  in
blanchet@52318
  1219
    fold_rev do_chunk chunks ([], [])
blanchet@52318
  1220
    |>> cons []
blanchet@52319
  1221
    ||> map nickname_of_thm
blanchet@52318
  1222
  end
blanchet@52318
  1223
blanchet@52319
  1224
fun attach_parents_to_facts _ [] = []
blanchet@52319
  1225
  | attach_parents_to_facts old_facts (facts as (_, th) :: _) =
blanchet@52319
  1226
    let
blanchet@52319
  1227
      fun do_facts _ [] = []
blanchet@52319
  1228
        | do_facts (_, parents) [fact] = [(parents, fact)]
blanchet@52319
  1229
        | do_facts (chunks, parents)
blanchet@52319
  1230
                   ((fact as (_, th)) :: (facts as (_, th') :: _)) =
blanchet@52319
  1231
          let
blanchet@52319
  1232
            val chunks = app_hd (cons th) chunks
blanchet@52319
  1233
            val chunks_and_parents' =
blanchet@58712
  1234
              if thm_less_eq (th, th') andalso thy_name_of_thm th = thy_name_of_thm th' then
blanchet@52319
  1235
                (chunks, [nickname_of_thm th])
blanchet@52319
  1236
              else
blanchet@52319
  1237
                chunks_and_parents_for chunks th'
blanchet@58348
  1238
          in
blanchet@58348
  1239
            (parents, fact) :: do_facts chunks_and_parents' facts
blanchet@58348
  1240
          end
blanchet@52319
  1241
    in
blanchet@52319
  1242
      old_facts @ facts
blanchet@52319
  1243
      |> do_facts (chunks_and_parents_for [[]] th)
blanchet@52319
  1244
      |> drop (length old_facts)
blanchet@52319
  1245
    end
blanchet@52314
  1246
blanchet@54232
  1247
fun maximal_wrt_graph G keys =
blanchet@54232
  1248
  let
blanchet@54232
  1249
    val tab = Symtab.empty |> fold (fn name => Symtab.default (name, ())) keys
blanchet@58348
  1250
blanchet@58354
  1251
    fun insert_new seen name = not (Symtab.defined seen name) ? insert (op =) name
blanchet@58348
  1252
blanchet@54232
  1253
    fun num_keys keys = Graph.Keys.fold (K (Integer.add 1)) keys 0
blanchet@58348
  1254
blanchet@54232
  1255
    fun find_maxes _ (maxs, []) = map snd maxs
blanchet@54232
  1256
      | find_maxes seen (maxs, new :: news) =
blanchet@58354
  1257
        find_maxes (seen |> num_keys (Graph.imm_succs G new) > 1 ? Symtab.default (new, ()))
blanchet@58354
  1258
          (if Symtab.defined tab new then
blanchet@58354
  1259
             let
blanchet@58354
  1260
               val newp = Graph.all_preds G [new]
blanchet@58354
  1261
               fun is_ancestor x yp = member (op =) yp x
blanchet@58354
  1262
               val maxs = maxs |> filter (fn (_, max) => not (is_ancestor max newp))
blanchet@58354
  1263
             in
blanchet@58354
  1264
               if exists (is_ancestor new o fst) maxs then (maxs, news)
blanchet@58354
  1265
               else ((newp, new) :: filter_out (fn (_, max) => is_ancestor max newp) maxs, news)
blanchet@58354
  1266
             end
blanchet@58354
  1267
           else
blanchet@58354
  1268
             (maxs, Graph.Keys.fold (insert_new seen) (Graph.imm_preds G new) news))
blanchet@58348
  1269
  in
blanchet@58348
  1270
    find_maxes Symtab.empty ([], Graph.maximals G)
blanchet@58348
  1271
  end
blanchet@54232
  1272
blanchet@58354
  1273
fun maximal_wrt_access_graph access_G facts =
blanchet@58355
  1274
  map (nickname_of_thm o snd) facts
blanchet@58354
  1275
  |> maximal_wrt_graph access_G
blanchet@54232
  1276
blanchet@54232
  1277
fun is_fact_in_graph access_G = can (Graph.get_node access_G) o nickname_of_thm
blanchet@54232
  1278
blanchet@54334
  1279
val chained_feature_factor = 0.5 (* FUDGE *)
blanchet@54334
  1280
val extra_feature_factor = 0.1 (* FUDGE *)
blanchet@58712
  1281
val num_extra_feature_facts = 0 (* FUDGE *) (* TODO: keep or eliminate? *)
blanchet@54232
  1282
blanchet@54232
  1283
(* FUDGE *)
blanchet@54232
  1284
fun weight_of_proximity_fact rank =
blanchet@54232
  1285
  Math.pow (1.3, 15.5 - 0.2 * Real.fromInt rank) + 15.0
blanchet@54232
  1286
blanchet@54277
  1287
fun weight_facts_smoothly facts =
blanchet@54232
  1288
  facts ~~ map weight_of_proximity_fact (0 upto length facts - 1)
blanchet@54232
  1289
blanchet@54277
  1290
(* FUDGE *)
blanchet@54277
  1291
fun steep_weight_of_fact rank =
blanchet@54277
  1292
  Math.pow (0.62, log2 (Real.fromInt (rank + 1)))
blanchet@54277
  1293
blanchet@54277
  1294
fun weight_facts_steeply facts =
blanchet@54277
  1295
  facts ~~ map steep_weight_of_fact (0 upto length facts - 1)
blanchet@54277
  1296
blanchet@58728
  1297
val max_proximity_facts = 100
blanchet@54232
  1298
blanchet@55512
  1299
fun find_mash_suggestions ctxt max_facts suggs facts chained raw_unknown =
blanchet@55512
  1300
  let
blanchet@55512
  1301
    val inter_fact = inter (eq_snd Thm.eq_thm_prop)
blanchet@55512
  1302
    val raw_mash = find_suggested_facts ctxt facts suggs
blanchet@55512
  1303
    val proximate = take max_proximity_facts facts
blanchet@55512
  1304
    val unknown_chained = inter_fact raw_unknown chained
blanchet@55512
  1305
    val unknown_proximate = inter_fact raw_unknown proximate
blanchet@55512
  1306
    val mess =
blanchet@55512
  1307
      [(0.9 (* FUDGE *), (map (rpair 1.0) unknown_chained, [])),
blanchet@55512
  1308
       (0.4 (* FUDGE *), (weight_facts_smoothly unknown_proximate, [])),
blanchet@55512
  1309
       (0.1 (* FUDGE *), (weight_facts_steeply raw_mash, raw_unknown))]
blanchet@58349
  1310
    val unknown = raw_unknown
blanchet@58349
  1311
      |> fold (subtract (eq_snd Thm.eq_thm_prop)) [unknown_chained, unknown_proximate]
blanchet@58348
  1312
  in
blanchet@58492
  1313
    (mesh_facts (eq_snd (gen_eq_thm ctxt)) max_facts mess, unknown)
blanchet@58348
  1314
  end
blanchet@54232
  1315
blanchet@54264
  1316
fun add_const_counts t =
blanchet@58348
  1317
  fold (fn s => Symtab.map_default (s, 0) (Integer.add 1)) (Term.add_const_names t [])
blanchet@54232
  1318
blanchet@58718
  1319
fun mash_suggested_facts ctxt ({debug, overlord, ...} : params) max_suggs hyp_ts concl_t facts =
blanchet@54232
  1320
  let
blanchet@54232
  1321
    val thy = Proof_Context.theory_of ctxt
blanchet@54696
  1322
    val thy_name = Context.theory_name thy
blanchet@58394
  1323
    val engine = the_mash_engine ()
blanchet@58394
  1324
blanchet@54278
  1325
    val facts = facts |> sort (crude_thm_ord o pairself snd o swap)
blanchet@54232
  1326
    val chained = facts |> filter (fn ((_, (scope, _)), _) => scope = Chained)
blanchet@54278
  1327
    val num_facts = length facts
blanchet@58402
  1328
blanchet@58438
  1329
    (* Weights appear to hurt kNN more than they help. *)
blanchet@58446
  1330
    val const_tab = Symtab.empty |> engine <> MaSh_SML_kNN
blanchet@58446
  1331
      ? fold (add_const_counts o prop_of o snd) facts
blanchet@55537
  1332
blanchet@54696
  1333
    fun fact_has_right_theory (_, th) =
blanchet@54696
  1334
      thy_name = Context.theory_name (theory_of_thm th)
blanchet@58337
  1335
blanchet@54278
  1336
    fun chained_or_extra_features_of factor (((_, stature), th), weight) =
blanchet@54278
  1337
      [prop_of th]
blanchet@58397
  1338
      |> features_of ctxt (theory_of_thm th) num_facts const_tab stature
blanchet@54278
  1339
      |> map (apsnd (fn r => weight * factor * r))
blanchet@55537
  1340
blanchet@58359
  1341
    fun query_args access_G =
blanchet@58359
  1342
      let
blanchet@58359
  1343
        val parents = maximal_wrt_access_graph access_G facts
blanchet@58359
  1344
        val hints = chained
blanchet@58359
  1345
          |> filter (is_fact_in_graph access_G o snd)
blanchet@58359
  1346
          |> map (nickname_of_thm o snd)
blanchet@58359
  1347
blanchet@58359
  1348
        val goal_feats =
blanchet@58397
  1349
          features_of ctxt thy num_facts const_tab (Local, General) (concl_t :: hyp_ts)
blanchet@58359
  1350
        val chained_feats = chained
blanchet@58359
  1351
          |> map (rpair 1.0)
blanchet@58359
  1352
          |> map (chained_or_extra_features_of chained_feature_factor)
blanchet@58359
  1353
          |> rpair [] |-> fold (union (eq_fst (op =)))
blanchet@58394
  1354
        val extra_feats =
blanchet@58437
  1355
          facts
blanchet@58437
  1356
          |> take (Int.max (0, num_extra_feature_facts - length chained))
blanchet@58437
  1357
          |> filter fact_has_right_theory
blanchet@58437
  1358
          |> weight_facts_steeply
blanchet@58437
  1359
          |> map (chained_or_extra_features_of extra_feature_factor)
blanchet@58437
  1360
          |> rpair [] |-> fold (union (eq_fst (op =)))
blanchet@58359
  1361
        val feats = fold (union (eq_fst (op =))) [chained_feats, extra_feats] goal_feats
blanchet@58359
  1362
          |> debug ? sort (Real.compare o swap o pairself snd)
blanchet@58359
  1363
      in
blanchet@58359
  1364
        (parents, hints, feats)
blanchet@58359
  1365
      end
blanchet@58359
  1366
blanchet@58720
  1367
    val ((access_G, ((num_facts, fact_tab), (num_feats, feat_tab)), ffds, freqs), py_suggs) =
blanchet@58720
  1368
      peek_state ctxt overlord (fn {access_G, xtabs, ffds, freqs, ...} =>
blanchet@58720
  1369
        ((access_G, xtabs, ffds, freqs),
blanchet@58713
  1370
         if Graph.is_empty access_G then
blanchet@58713
  1371
           (trace_msg ctxt (K "Nothing has been learned yet"); [])
blanchet@58713
  1372
         else if engine = MaSh_Py then
blanchet@58713
  1373
           let val (parents, hints, feats) = query_args access_G in
blanchet@58718
  1374
             MaSh_Py.query ctxt overlord max_suggs ([], hints, parents, feats)
blanchet@58713
  1375
             |> map fst
blanchet@58713
  1376
           end
blanchet@58713
  1377
         else
blanchet@58713
  1378
           []))
blanchet@58359
  1379
blanchet@58359
  1380
    val sml_suggs =
blanchet@58371
  1381
      if engine = MaSh_Py then
blanchet@58371
  1382
        []
blanchet@58371
  1383
      else
blanchet@58445
  1384
        let
blanchet@58718
  1385
          val (parents, hints, goal_feats0) = query_args access_G
blanchet@58718
  1386
          val goal_feats = map fst goal_feats0
blanchet@58718
  1387
          val visible_facts = map_filter (Symtab.lookup fact_tab) (Graph.all_preds access_G parents)
blanchet@58445
  1388
        in
blanchet@58718
  1389
          if engine = MaSh_SML_kNN_Ext orelse engine = MaSh_SML_NB_Ext then
blanchet@58718
  1390
            let
blanchet@58718
  1391
              val learns =
blanchet@58718
  1392
                Graph.schedule (fn _ => fn (fact, (_, feats, deps)) => (fact, feats, deps)) access_G
blanchet@58718
  1393
            in
blanchet@58718
  1394
              MaSh_SML.query_external ctxt engine max_suggs learns goal_feats
blanchet@58718
  1395
            end
blanchet@58718
  1396
          else
blanchet@58718
  1397
            let
blanchet@58718
  1398
              val int_goal_feats = map_filter (Symtab.lookup feat_tab) goal_feats
blanchet@58718
  1399
            in
blanchet@58720
  1400
              MaSh_SML.query_internal ctxt engine num_facts num_feats ffds freqs visible_facts
blanchet@58720
  1401
                max_suggs goal_feats int_goal_feats
blanchet@58718
  1402
            end
blanchet@58359
  1403
        end
blanchet@58359
  1404
blanchet@58347
  1405
    val unknown = filter_out (is_fact_in_graph access_G o snd) facts
blanchet@54232
  1406
  in
blanchet@58718
  1407
    find_mash_suggestions ctxt max_suggs (py_suggs @ sml_suggs) facts chained unknown
blanchet@54232
  1408
    |> pairself (map fact_of_raw_fact)
blanchet@54232
  1409
  end
blanchet@54232
  1410
blanchet@58723
  1411
fun learn_wrt_access_graph ctxt (name, parents, feats, deps) (access_G, (fact_xtab, feat_xtab)) =
blanchet@54232
  1412
  let
blanchet@58713
  1413
    fun maybe_learn_from from (accum as (parents, access_G)) =
blanchet@58355
  1414
      try_graph ctxt "updating graph" accum (fn () =>
blanchet@58713
  1415
        (from :: parents, Graph.add_edge_acyclic (from, name) access_G))
blanchet@58713
  1416
blanchet@58713
  1417
    val access_G = access_G |> Graph.default_node (name, (Isar_Proof, feats, deps))
blanchet@58713
  1418
    val (parents, access_G) = ([], access_G) |> fold maybe_learn_from parents
blanchet@58713
  1419
    val (deps, _) = ([], access_G) |> fold maybe_learn_from deps
blanchet@58713
  1420
blanchet@58725
  1421
    val fact_xtab = add_to_xtab name fact_xtab
blanchet@58713
  1422
    val feat_xtab = fold maybe_add_to_xtab feats feat_xtab
blanchet@58348
  1423
  in
blanchet@58723
  1424
    ((name, parents, feats, deps), (access_G, (fact_xtab, feat_xtab)))
blanchet@58348
  1425
  end
blanchet@54232
  1426
blanchet@58723
  1427
fun relearn_wrt_access_graph ctxt (name, deps) access_G =
blanchet@54232
  1428
  let
blanchet@58713
  1429
    fun maybe_relearn_from from (accum as (parents, access_G)) =
blanchet@54232
  1430
      try_graph ctxt "updating graph" accum (fn () =>
blanchet@58713
  1431
        (from :: parents, Graph.add_edge_acyclic (from, name) access_G))
blanchet@58713
  1432
    val access_G =
blanchet@58713
  1433
      access_G |> Graph.map_node name (fn (_, feats, _) => (Automatic_Proof, feats, deps))
blanchet@58713
  1434
    val (deps, _) = ([], access_G) |> fold maybe_relearn_from deps
blanchet@58348
  1435
  in
blanchet@58723
  1436
    ((name, deps), access_G)
blanchet@58348
  1437
  end
blanchet@54232
  1438
blanchet@54232
  1439
fun flop_wrt_access_graph name =
blanchet@58347
  1440
  Graph.map_node name (fn (_, feats, deps) => (Isar_Proof_wegen_Prover_Flop, feats, deps))
blanchet@54232
  1441
blanchet@58615
  1442
val learn_timeout_slack = 20.0
blanchet@54232
  1443
blanchet@54232
  1444
fun launch_thread timeout task =
blanchet@54232
  1445
  let
blanchet@54232
  1446
    val hard_timeout = time_mult learn_timeout_slack timeout
blanchet@54232
  1447
    val birth_time = Time.now ()
blanchet@54232
  1448
    val death_time = Time.+ (birth_time, hard_timeout)
blanchet@54232
  1449
    val desc = ("Machine learner for Sledgehammer", "")
blanchet@58348
  1450
  in
blanchet@58348
  1451
    Async_Manager.thread MaShN birth_time death_time desc task
blanchet@58348
  1452
  end
blanchet@54232
  1453
blanchet@58355
  1454
fun learned_proof_name () =
blanchet@58355
  1455
  Date.fmt ".%Y%m%d.%H%M%S." (Date.fromTimeLocal (Time.now ())) ^ serial_string ()
blanchet@58355
  1456
blanchet@55876
  1457
fun mash_learn_proof ctxt ({overlord, timeout, ...} : params) t facts used_ths =
blanchet@54956
  1458
  if is_mash_enabled () then
blanchet@56158
  1459
    launch_thread timeout (fn () =>
blanchet@58348
  1460
      let
blanchet@58348
  1461
        val thy = Proof_Context.theory_of ctxt
blanchet@58700
  1462
        val feats = map fst (features_of ctxt thy 0 Symtab.empty (Local, General) [t])
blanchet@58348
  1463
      in
blanchet@58713
  1464
        map_state ctxt overlord
blanchet@58722
  1465
          (fn state as {access_G, xtabs as ((num_facts0, _), _), ffds, freqs, dirty_facts} =>
blanchet@58713
  1466
             let
blanchet@58713
  1467
               val parents = maximal_wrt_access_graph access_G facts
blanchet@58713
  1468
               val deps = used_ths
blanchet@58713
  1469
                 |> filter (is_fact_in_graph access_G)
blanchet@58713
  1470
                 |> map nickname_of_thm
blanchet@58713
  1471
             in
blanchet@58713
  1472
               if the_mash_engine () = MaSh_Py then
blanchet@58713
  1473
                 (MaSh_Py.learn ctxt overlord true [("", parents, feats, deps)]; state)
blanchet@58713
  1474
               else
blanchet@58713
  1475
                 let
blanchet@58713
  1476
                   val name = learned_proof_name ()
blanchet@58722
  1477
                   val (access_G', xtabs', rev_learns) =
blanchet@58721
  1478
                     add_node Automatic_Proof name parents feats deps (access_G, xtabs, [])
blanchet@58720
  1479
blanchet@58722
  1480
                   val (ffds', freqs') =
blanchet@58722
  1481
                     recompute_ffds_freqs_from_learns (rev rev_learns) xtabs' num_facts0 ffds freqs
blanchet@58713
  1482
                 in
blanchet@58720
  1483
                   {access_G = access_G', xtabs = xtabs', ffds = ffds', freqs = freqs',
blanchet@58713
  1484
                    dirty_facts = Option.map (cons name) dirty_facts}
blanchet@58713
  1485
                 end
blanchet@58713
  1486
             end);
blanchet@58348
  1487
        (true, "")
blanchet@58348
  1488
      end)
blanchet@54956
  1489
  else
blanchet@54956
  1490
    ()
blanchet@54232
  1491
blanchet@58695
  1492
fun sendback sub = Active.sendback_markup [Markup.padding_command] (sledgehammerN ^ " " ^ sub)
blanchet@49407
  1493
blanchet@49407
  1494
val commit_timeout = seconds 30.0
blanchet@49347
  1495
blanchet@51500
  1496
(* The timeout is understood in a very relaxed fashion. *)
blanchet@56158
  1497
fun mash_learn_facts ctxt (params as {debug, verbose, overlord, ...}) prover save auto_level
blanchet@56158
  1498
    run_prover learn_timeout facts =
blanchet@49319
  1499
  let
blanchet@49333
  1500
    val timer = Timer.startRealTimer ()
blanchet@56158
  1501
    fun next_commit_time () = Time.+ (Timer.checkRealTimer timer, commit_timeout)
blanchet@58347
  1502
blanchet@58371
  1503
    val engine = the_mash_engine ()
blanchet@58714
  1504
    val {access_G, ...} = peek_state ctxt overlord I
blanchet@54232
  1505
    val is_in_access_G = is_fact_in_graph access_G o snd
blanchet@52314
  1506
    val no_new_facts = forall is_in_access_G facts
blanchet@49323
  1507
  in
blanchet@52314
  1508
    if no_new_facts andalso not run_prover then
blanchet@49419
  1509
      if auto_level < 2 then
blanchet@58347
  1510
        "No new " ^ (if run_prover then "automatic" else "Isar") ^ " proofs to learn." ^
blanchet@51499
  1511
        (if auto_level = 0 andalso not run_prover then
blanchet@58347
  1512
           "\n\nHint: Try " ^ sendback learn_proverN ^ " to learn from an automatic prover."
blanchet@49419
  1513
         else
blanchet@49419
  1514
           "")
blanchet@49407
  1515
      else
blanchet@49407
  1516
        ""
blanchet@49323
  1517
    else
blanchet@49319
  1518
      let
blanchet@51750
  1519
        val name_tabs = build_name_tables nickname_of_thm facts
blanchet@58347
  1520
blanchet@49454
  1521
        fun deps_of status th =
blanchet@58359
  1522
          if status = Non_Rec_Def orelse status = Rec_Def then
blanchet@49454
  1523
            SOME []
blanchet@51499
  1524
          else if run_prover then
blanchet@58347
  1525
            prover_dependencies_of ctxt params prover auto_level facts name_tabs th
blanchet@58347
  1526
            |> (fn (false, _) => NONE | (true, deps) => trim_dependencies deps)
blanchet@49419
  1527
          else
blanchet@51750
  1528
            isar_dependencies_of name_tabs th
blanchet@58347
  1529
blanchet@49684
  1530
        fun do_commit [] [] [] state = state
blanchet@58724
  1531
          | do_commit learns relearns flops
blanchet@58724
  1532
              {access_G, xtabs as ((num_facts0, _), _), ffds, freqs, dirty_facts} =
blanchet@49407
  1533
            let
blanchet@58720
  1534
              val was_empty = Graph.is_empty access_G
blanchet@58720
  1535
blanchet@58716
  1536
              val (learns, (access_G, xtabs)) =
blanchet@58723
  1537
                fold_map (learn_wrt_access_graph ctxt) learns (access_G, xtabs)
blanchet@58713
  1538
              val (relearns, access_G) =
blanchet@58723
  1539
                fold_map (relearn_wrt_access_graph ctxt) relearns access_G
blanchet@58713
  1540
blanchet@51625
  1541
              val access_G = access_G |> fold flop_wrt_access_graph flops
blanchet@58707
  1542
              val dirty_facts =
blanchet@58707
  1543
                (case (was_empty, dirty_facts) of
blanchet@58404
  1544
                  (false, SOME names) => SOME (map #1 learns @ map #1 relearns @ names)
blanchet@56628
  1545
                | _ => NONE)
blanchet@58720
  1546
blanchet@58724
  1547
              val (ffds', freqs') =
blanchet@58724
  1548
                if null relearns then
blanchet@58724
  1549
                  recompute_ffds_freqs_from_learns
blanchet@58724
  1550
                    (map (fn (name, _, feats, deps) => (name, feats, deps)) learns) xtabs num_facts0
blanchet@58724
  1551
                    ffds freqs
blanchet@58724
  1552
                else
blanchet@58724
  1553
                  recompute_ffds_freqs_from_access_G access_G xtabs
blanchet@49419
  1554
            in
blanchet@58371
  1555
              if engine = MaSh_Py then
blanchet@58723
  1556
                (MaSh_Py.learn ctxt overlord (save andalso null relearns) learns;
blanchet@58371
  1557
                 MaSh_Py.relearn ctxt overlord save relearns)
blanchet@58349
  1558
              else
blanchet@58371
  1559
                ();
blanchet@58720
  1560
              {access_G = access_G, xtabs = xtabs, ffds = ffds', freqs = freqs',
blanchet@58720
  1561
               dirty_facts = dirty_facts}
blanchet@49419
  1562
            end
blanchet@58347
  1563
blanchet@51646
  1564
        fun commit last learns relearns flops =
blanchet@58359
  1565
          (if debug andalso auto_level = 0 then Output.urgent_message "Committing..." else ();
blanchet@54285
  1566
           map_state ctxt overlord (do_commit (rev learns) relearns flops);
blanchet@49419
  1567
           if not last andalso auto_level = 0 then
blanchet@51646
  1568
             let val num_proofs = length learns + length relearns in
blanchet@58359
  1569
               Output.urgent_message ("Learned " ^ string_of_int num_proofs ^ " " ^
blanchet@58359
  1570
                 (if run_prover then "automatic" else "Isar") ^ " proof" ^
blanchet@58359
  1571
                 plural_s num_proofs ^ " in the last " ^ string_of_time commit_timeout ^ ".")
blanchet@49407
  1572
             end
blanchet@49407
  1573
           else
blanchet@49407
  1574
             ())
blanchet@58347
  1575
blanchet@52314
  1576
        fun learn_new_fact _ (accum as (_, (_, _, true))) = accum
blanchet@52314
  1577
          | learn_new_fact (parents, ((_, stature as (_, status)), th))
blanchet@58727
  1578
              (learns, (num_nontrivial, next_commit, _)) =
blanchet@49333
  1579
            let
blanchet@51639
  1580
              val name = nickname_of_thm th
blanchet@58700
  1581
              val feats =
blanchet@58700
  1582
                map fst (features_of ctxt (theory_of_thm th) 0 Symtab.empty stature [prop_of th])
blanchet@49454
  1583
              val deps = deps_of status th |> these
blanchet@58727
  1584
              val num_nontrivial = num_nontrivial |> not (null deps) ? Integer.add 1
blanchet@51646
  1585
              val learns = (name, parents, feats, deps) :: learns
blanchet@51646
  1586
              val (learns, next_commit) =
blanchet@49407
  1587
                if Time.> (Timer.checkRealTimer timer, next_commit) then
blanchet@51646
  1588
                  (commit false learns [] []; ([], next_commit_time ()))
blanchet@49407
  1589
                else
blanchet@51646
  1590
                  (learns, next_commit)
blanchet@56158
  1591
              val timed_out = Time.> (Timer.checkRealTimer timer, learn_timeout)
blanchet@58347
  1592
            in
blanchet@58727
  1593
              (learns, (num_nontrivial, next_commit, timed_out))
blanchet@58347
  1594
            end
blanchet@58347
  1595
blanchet@58727
  1596
        val (num_new_facts, num_nontrivial) =
blanchet@52314
  1597
          if no_new_facts then
blanchet@58727
  1598
            (0, 0)
blanchet@49419
  1599
          else
blanchet@49419
  1600
            let
blanchet@58347
  1601
              val new_facts = facts
blanchet@58347
  1602
                |> sort (crude_thm_ord o pairself snd)
blanchet@58347
  1603
                |> attach_parents_to_facts []
blanchet@58347
  1604
                |> filter_out (is_in_access_G o snd)
blanchet@58727
  1605
              val (learns, (num_nontrivial, _, _)) =
blanchet@52314
  1606
                ([], (0, next_commit_time (), false))
blanchet@52319
  1607
                |> fold learn_new_fact new_facts
blanchet@58347
  1608
            in
blanchet@58727
  1609
              commit true learns [] []; (length new_facts, num_nontrivial)
blanchet@58347
  1610
            end
blanchet@58347
  1611
blanchet@49419
  1612
        fun relearn_old_fact _ (accum as (_, (_, _, true))) = accum
blanchet@58727
  1613
          | relearn_old_fact ((_, (_, status)), th)
blanchet@58727
  1614
              ((relearns, flops), (num_nontrivial, next_commit, _)) =
blanchet@49419
  1615
            let
blanchet@51639
  1616
              val name = nickname_of_thm th
blanchet@58727
  1617
              val (num_nontrivial, relearns, flops) =
blanchet@56628
  1618
                (case deps_of status th of
blanchet@58727
  1619
                  SOME deps => (num_nontrivial + 1, (name, deps) :: relearns, flops)
blanchet@58727
  1620
                | NONE => (num_nontrivial, relearns, name :: flops))
blanchet@51646
  1621
              val (relearns, flops, next_commit) =
blanchet@49419
  1622
                if Time.> (Timer.checkRealTimer timer, next_commit) then
blanchet@58347
  1623
                  (commit false [] relearns flops; ([], [], next_commit_time ()))
blanchet@49419
  1624
                else
blanchet@51646
  1625
                  (relearns, flops, next_commit)
blanchet@56158
  1626
              val timed_out = Time.> (Timer.checkRealTimer timer, learn_timeout)
blanchet@58347
  1627
            in
blanchet@58727
  1628
              ((relearns, flops), (num_nontrivial, next_commit, timed_out))
blanchet@58347
  1629
            end
blanchet@58347
  1630
blanchet@58727
  1631
        val num_nontrivial =
blanchet@52314
  1632
          if not run_prover then
blanchet@58727
  1633
            num_nontrivial
blanchet@49419
  1634
          else
blanchet@49419
  1635
            let
blanchet@49683
  1636
              val max_isar = 1000 * max_dependencies
blanchet@58347
  1637
blanchet@58347
  1638
              fun priority_of th =
blanchet@58404
  1639
                random_range 0 max_isar +
blanchet@58404
  1640
                (case try (Graph.get_node access_G) (nickname_of_thm th) of
blanchet@58404
  1641
                  SOME (Isar_Proof, _, deps) => ~100 * length deps
blanchet@58404
  1642
                | SOME (Automatic_Proof, _, _) => 2 * max_isar
blanchet@58404
  1643
                | SOME (Isar_Proof_wegen_Prover_Flop, _, _) => max_isar
blanchet@58404
  1644
                | NONE => 0)
blanchet@58347
  1645
blanchet@58347
  1646
              val old_facts = facts
blanchet@58347
  1647
                |> filter is_in_access_G
blanchet@58347
  1648
                |> map (`(priority_of o snd))
blanchet@58347
  1649
                |> sort (int_ord o pairself fst)
blanchet@58347
  1650
                |> map snd
blanchet@58727
  1651
              val ((relearns, flops), (num_nontrivial, _, _)) =
blanchet@58727
  1652
                (([], []), (num_nontrivial, next_commit_time (), false))
blanchet@49419
  1653
                |> fold relearn_old_fact old_facts
blanchet@58347
  1654
            in
blanchet@58727
  1655
              commit true [] relearns flops; num_nontrivial
blanchet@58347
  1656
            end
blanchet@49333
  1657
      in
blanchet@49419
  1658
        if verbose orelse auto_level < 2 then
blanchet@58727
  1659
          "Learned " ^ string_of_int num_new_facts ^ " fact" ^ plural_s num_new_facts ^ " and " ^
blanchet@58727
  1660
          string_of_int num_nontrivial ^ " nontrivial " ^
blanchet@58727
  1661
          (if run_prover then "automatic and " else "") ^ "Isar proof" ^ plural_s num_nontrivial ^
blanchet@58348
  1662
          (if verbose then " in " ^ string_of_time (Timer.checkRealTimer timer) else "") ^ "."
blanchet@49334
  1663
        else
blanchet@49334
  1664
          ""
blanchet@49333
  1665
      end
blanchet@49323
  1666
  end
blanchet@49319
  1667
blanchet@55575
  1668
fun mash_learn ctxt (params as {provers, timeout, ...}) fact_override chained run_prover =
blanchet@49331
  1669
  let
blanchet@49411
  1670
    val css = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
blanchet@49410
  1671
    val ctxt = ctxt |> Config.put instantiate_inducts false
blanchet@58348
  1672
    val facts = nearly_all_facts ctxt false fact_override Symtab.empty css chained [] @{prop True}
blanchet@55567
  1673
      |> sort (crude_thm_ord o pairself snd o swap)
blanchet@49419
  1674
    val num_facts = length facts
blanchet@49419
  1675
    val prover = hd provers
blanchet@58348
  1676
blanchet@51499
  1677
    fun learn auto_level run_prover =
blanchet@56158
  1678
      mash_learn_facts ctxt params prover true auto_level run_prover one_year facts
blanchet@49419
  1679
      |> Output.urgent_message
blanchet@49331
  1680
  in
blanchet@51499
  1681
    if run_prover then
blanchet@58359
  1682
      (Output.urgent_message ("MaShing through " ^ string_of_int num_facts ^ " fact" ^
blanchet@58359
  1683
         plural_s num_facts ^ " for automatic proofs (" ^ quote prover ^ " timeout: " ^
blanchet@58359
  1684
         string_of_time timeout ^ ").\n\nCollecting Isar proofs first...");
blanchet@51355
  1685
       learn 1 false;
blanchet@58359
  1686
       Output.urgent_message "Now collecting automatic proofs. This may take several hours. You \
blanchet@58359
  1687
         \can safely stop the learning process at any point.";
blanchet@51355
  1688
       learn 0 true)
blanchet@51355
  1689
    else
blanchet@58348
  1690
      (Output.urgent_message ("MaShing through " ^ string_of_int num_facts ^ " fact" ^
blanchet@58348
  1691
         plural_s num_facts ^ " for Isar proofs...");
blanchet@51355
  1692
       learn 0 false)
blanchet@49331
  1693
  end
blanchet@49264
  1694
blanchet@54285
  1695
fun mash_can_suggest_facts ctxt overlord =
blanchet@54285
  1696
  not (Graph.is_empty (#access_G (peek_state ctxt overlord I)))
blanchet@51326
  1697
blanchet@58616
  1698
(* Generate more suggestions than requested, because some might be thrown out later for various
blanchet@58616
  1699
   reasons (e.g., duplicates). *)
blanchet@58616
  1700
fun generous_max_suggestions max_facts = 3 * max_facts div 2 + 25
blanchet@51398
  1701
blanchet@51829
  1702
val mepo_weight = 0.5
blanchet@51829
  1703
val mash_weight = 0.5
blanchet@51829
  1704
blanchet@54289
  1705
val max_facts_to_learn_before_query = 100
blanchet@54289
  1706
blanchet@58347
  1707
(* The threshold should be large enough so that MaSh does not get activated for Auto Sledgehammer
blanchet@58347
  1708
   and Try. *)
blanchet@49333
  1709
val min_secs_for_learning = 15
blanchet@49333
  1710
blanchet@58703
  1711
fun relevant_facts ctxt (params as {overlord, learn, fact_filter, timeout, ...}) prover max_facts
blanchet@58703
  1712
    ({add, only, ...} : fact_override) hyp_ts concl_t facts =
blanchet@49329
  1713
  if not (subset (op =) (the_list fact_filter, fact_filters)) then
blanchet@49329
  1714
    error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".")
blanchet@49329
  1715
  else if only then
blanchet@58492
  1716
    [("", map fact_of_raw_fact facts)]
blanchet@49336
  1717
  else if max_facts <= 0 orelse null facts then
blanchet@52192
  1718
    [("", [])]
blanchet@49303
  1719
  else
blanchet@49303
  1720
    let
blanchet@58726
  1721
      fun maybe_launch_thread min_num_facts_to_learn =
blanchet@58615
  1722
        if not (Async_Manager.has_running_threads MaShN) andalso
blanchet@56158
  1723
           Time.toSeconds timeout >= min_secs_for_learning then
blanchet@56158
  1724
          let val timeout = time_mult learn_timeout_slack timeout in
blanchet@58726
  1725
            Output.urgent_message ("Started MaShing through at least " ^
blanchet@58726
  1726
              string_of_int min_num_facts_to_learn ^ " fact" ^ plural_s min_num_facts_to_learn ^
blanchet@58728
  1727
              " in the background.");
blanchet@56158
  1728
            launch_thread timeout
blanchet@56158
  1729
              (fn () => (true, mash_learn_facts ctxt params prover true 2 false timeout facts))
blanchet@49334
  1730
          end
blanchet@49333
  1731
        else
blanchet@49333
  1732
          ()
blanchet@58360
  1733
blanchet@54289
  1734
      fun maybe_learn () =
blanchet@54956
  1735
        if is_mash_enabled () andalso learn then
blanchet@54289
  1736
          let
blanchet@58719
  1737
            val {access_G, xtabs = ((num_facts0, _), _), ...} = peek_state ctxt overlord I
blanchet@54289
  1738
            val is_in_access_G = is_fact_in_graph access_G o snd
blanchet@58726
  1739
            val min_num_facts_to_learn = length facts - num_facts0
blanchet@54289
  1740
          in
blanchet@58726
  1741
            if min_num_facts_to_learn <= max_facts_to_learn_before_query then
blanchet@56628
  1742
              (case length (filter_out is_in_access_G facts) of
blanchet@55149
  1743
                0 => false
blanchet@55149
  1744
              | num_facts_to_learn =>
blanchet@55149
  1745
                if num_facts_to_learn <= max_facts_to_learn_before_query then
blanchet@55149
  1746
                  (mash_learn_facts ctxt params prover false 2 false timeout facts
blanchet@55149
  1747
                   |> (fn "" => () | s => Output.urgent_message (MaShN ^ ": " ^ s));
blanchet@55149
  1748
                   true)
blanchet@55149
  1749
                else
blanchet@58726
  1750
                  (maybe_launch_thread num_facts_to_learn; false))
blanchet@54289
  1751
            else
blanchet@58726
  1752
              (maybe_launch_thread min_num_facts_to_learn; false)
blanchet@54289
  1753
          end
blanchet@54289
  1754
        else
blanchet@54289
  1755
          false
blanchet@58360
  1756
blanchet@54289
  1757
      val (save, effective_fact_filter) =
blanchet@56628
  1758
        (case fact_filter of
blanchet@54289
  1759
          SOME ff => (ff <> mepoN andalso maybe_learn (), ff)
blanchet@49333
  1760
        | NONE =>
blanchet@52206
  1761
          if is_mash_enabled () then
blanchet@58449
  1762
            (maybe_learn (), if mash_can_suggest_facts ctxt overlord then meshN else mepoN)
blanchet@49422
  1763
          else
blanchet@56628
  1764
            (false, mepoN))
blanchet@55595
  1765
blanchet@55595
  1766
      val unique_facts = drop_duplicate_facts facts
blanchet@49303
  1767
      val add_ths = Attrib.eval_thms ctxt add
blanchet@55595
  1768
blanchet@52186
  1769
      fun in_add (_, th) = member Thm.eq_thm_prop add_ths th
blanchet@58360
  1770
blanchet@52185
  1771
      fun add_and_take accepts =
blanchet@52185
  1772
        (case add_ths of
blanchet@52185
  1773
           [] => accepts
blanchet@58492
  1774
         | _ =>
blanchet@58492
  1775
           (unique_facts |> filter in_add |> map fact_of_raw_fact) @ (accepts |> filter_out in_add))
blanchet@49308
  1776
        |> take max_facts
blanchet@58360
  1777
blanchet@49421
  1778
      fun mepo () =
blanchet@55595
  1779
        (mepo_suggested_facts ctxt params max_facts NONE hyp_ts concl_t unique_facts
blanchet@55543
  1780
         |> weight_facts_steeply, [])
blanchet@58360
  1781
blanchet@49329
  1782
      fun mash () =
blanchet@58450
  1783
        mash_suggested_facts ctxt params (generous_max_suggestions max_facts) hyp_ts concl_t facts
blanchet@54277
  1784
        |>> weight_facts_steeply
blanchet@58360
  1785
blanchet@49329
  1786
      val mess =
blanchet@52185
  1787
        (* the order is important for the "case" expression below *)
blanchet@55543
  1788
        [] |> effective_fact_filter <> mepoN ? cons (mash_weight, mash)
blanchet@55543
  1789
           |> effective_fact_filter <> mashN ? cons (mepo_weight, mepo)
blanchet@55543
  1790
           |> Par_List.map (apsnd (fn f => f ()))
blanchet@58492
  1791
      val mesh = mesh_facts (eq_snd (gen_eq_thm ctxt)) max_facts mess |> add_and_take
blanchet@49303
  1792
    in
blanchet@58371
  1793
      if the_mash_engine () = MaSh_Py andalso save then MaSh_Py.save ctxt overlord else ();
blanchet@56628
  1794
      (case (fact_filter, mess) of
blanchet@52206
  1795
        (NONE, [(_, (mepo, _)), (_, (mash, _))]) =>
blanchet@52192
  1796
        [(meshN, mesh), (mepoN, mepo |> map fst |> add_and_take),
blanchet@52192
  1797
         (mashN, mash |> map fst |> add_and_take)]
blanchet@56628
  1798
      | _ => [(effective_fact_filter, mesh)])
blanchet@49303
  1799
    end
blanchet@49303
  1800
blanchet@54285
  1801
fun kill_learners ctxt ({overlord, ...} : params) =
blanchet@58349
  1802
  (Async_Manager.kill_threads MaShN "learner";
blanchet@58371
  1803
   if the_mash_engine () = MaSh_Py then MaSh_Py.shutdown ctxt overlord else ())
blanchet@58347
  1804
blanchet@49334
  1805
fun running_learners () = Async_Manager.running_threads MaShN "learner"
blanchet@49334
  1806
blanchet@49263
  1807
end;