src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML
author Walther Neuper <walther.neuper@jku.at>
Tue, 03 Sep 2019 16:10:31 +0200
changeset 59606 c3925099d59f
parent 59451 71b442e82416
child 60065 46266dc209cd
permissions -rw-r--r--
\----- start update Isabelle2018 --> Isabelle2019
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@51647
    35
blanchet@58874
    36
  datatype mash_algorithm =
blanchet@58800
    37
    MaSh_NB
blanchet@58801
    38
  | MaSh_kNN
blanchet@58801
    39
  | MaSh_NB_kNN
blanchet@58800
    40
  | MaSh_NB_Ext
blanchet@58773
    41
  | MaSh_kNN_Ext
blanchet@58448
    42
blanchet@58462
    43
  val is_mash_enabled : unit -> bool
blanchet@58874
    44
  val the_mash_algorithm : unit -> mash_algorithm
blanchet@58899
    45
  val str_of_mash_algorithm : mash_algorithm -> string
blanchet@58462
    46
wneuper@59324
    47
  val mesh_facts : ('a list -> 'a list) -> ('a * 'a -> bool) -> int ->
wneuper@59324
    48
    (real * (('a * real) list * 'a list)) list -> 'a list
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
wneuper@59324
    51
  val crude_thm_ord : Proof.context -> thm * thm -> order
blanchet@52319
    52
  val thm_less : thm * thm -> bool
blanchet@49266
    53
  val goal_of_thm : theory -> thm -> thm
blanchet@58348
    54
  val run_prover_for_mash : Proof.context -> params -> string -> string -> fact list -> thm ->
blanchet@58348
    55
    prover_result
wneuper@59324
    56
  val features_of : Proof.context -> string -> stature -> term list -> string list
blanchet@52314
    57
  val trim_dependencies : string list -> string list option
blanchet@58648
    58
  val isar_dependencies_of : string Symtab.table * string Symtab.table -> thm -> string list option
blanchet@58348
    59
  val prover_dependencies_of : Proof.context -> params -> string -> int -> raw_fact list ->
blanchet@58348
    60
    string Symtab.table * string Symtab.table -> thm -> bool * string list
blanchet@58348
    61
  val attach_parents_to_facts : ('a * thm) list -> ('a * thm) list ->
blanchet@58348
    62
    (string list * ('a * thm)) list
blanchet@54277
    63
  val num_extra_feature_facts : int
blanchet@54278
    64
  val extra_feature_factor : real
blanchet@54277
    65
  val weight_facts_smoothly : 'a list -> ('a * real) list
blanchet@54277
    66
  val weight_facts_steeply : 'a list -> ('a * real) list
blanchet@58774
    67
  val find_mash_suggestions : Proof.context -> int -> string list -> ('a * thm) list ->
blanchet@58774
    68
    ('a * thm) list -> ('a * thm) list -> ('a * thm) list * ('a * thm) list
wneuper@59324
    69
  val mash_suggested_facts : Proof.context -> string -> params -> int -> term list -> term ->
blanchet@58749
    70
    raw_fact list -> fact list * fact list
blanchet@58773
    71
wneuper@59324
    72
  val mash_unlearn : Proof.context -> unit
walther@59606
    73
  val mash_learn_proof : Proof.context -> params -> term -> thm list -> unit
blanchet@58773
    74
  val mash_learn_facts : Proof.context -> params -> string -> int -> bool -> Time.time ->
blanchet@58462
    75
    raw_fact list -> string
blanchet@58348
    76
  val mash_learn : Proof.context -> params -> fact_override -> thm list -> bool -> unit
blanchet@58774
    77
  val mash_can_suggest_facts : Proof.context -> bool
wneuper@59324
    78
  val mash_can_suggest_facts_fast : Proof.context -> bool
blanchet@54956
    79
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@49263
    85
end;
blanchet@49263
    86
blanchet@49396
    87
structure Sledgehammer_MaSh : SLEDGEHAMMER_MASH =
blanchet@49263
    88
struct
blanchet@49264
    89
blanchet@49266
    90
open ATP_Util
blanchet@49266
    91
open ATP_Problem_Generate
blanchet@49266
    92
open Sledgehammer_Util
blanchet@49266
    93
open Sledgehammer_Fact
blanchet@56543
    94
open Sledgehammer_Prover
blanchet@56544
    95
open Sledgehammer_Prover_Minimize
blanchet@49396
    96
open Sledgehammer_MePo
blanchet@49266
    97
wneuper@59324
    98
val anonymous_proof_prefix = "."
wneuper@59324
    99
walther@59606
   100
val trace = Attrib.setup_config_bool \<^binding>\<open>sledgehammer_mash_trace\<close> (K false)
walther@59606
   101
val duplicates = Attrib.setup_config_bool \<^binding>\<open>sledgehammer_fact_duplicates\<close> (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@58894
   132
fun state_file () = Path.expand (Path.explode "$ISABELLE_HOME_USER/mash_state")
blanchet@58894
   133
val remove_state_file = try File.rm o state_file
blanchet@49345
   134
blanchet@58874
   135
datatype mash_algorithm =
blanchet@58800
   136
  MaSh_NB
blanchet@58801
   137
| MaSh_kNN
blanchet@58801
   138
| MaSh_NB_kNN
blanchet@58800
   139
| MaSh_NB_Ext
blanchet@58773
   140
| MaSh_kNN_Ext
blanchet@58620
   141
blanchet@58874
   142
fun mash_algorithm () =
walther@59606
   143
  (case Options.default_string \<^system_option>\<open>MaSh\<close> of
wneuper@59324
   144
    "yes" => SOME MaSh_NB_kNN
wneuper@59324
   145
  | "sml" => SOME MaSh_NB_kNN
wneuper@59324
   146
  | "nb" => SOME MaSh_NB
wneuper@59324
   147
  | "knn" => SOME MaSh_kNN
wneuper@59324
   148
  | "nb_knn" => SOME MaSh_NB_kNN
wneuper@59324
   149
  | "nb_ext" => SOME MaSh_NB_Ext
wneuper@59324
   150
  | "knn_ext" => SOME MaSh_kNN_Ext
wneuper@59324
   151
  | "none" => NONE
wneuper@59324
   152
  | "" => NONE
wneuper@59324
   153
  | algorithm => (warning ("Unknown MaSh algorithm: " ^ quote algorithm); NONE))
blanchet@58360
   154
blanchet@58874
   155
val is_mash_enabled = is_some o mash_algorithm
blanchet@58874
   156
val the_mash_algorithm = the_default MaSh_NB_kNN o mash_algorithm
blanchet@58360
   157
blanchet@58899
   158
fun str_of_mash_algorithm MaSh_NB = "nb"
blanchet@58899
   159
  | str_of_mash_algorithm MaSh_kNN = "knn"
blanchet@58899
   160
  | str_of_mash_algorithm MaSh_NB_kNN = "nb_knn"
blanchet@58899
   161
  | str_of_mash_algorithm MaSh_NB_Ext = "nb_ext"
blanchet@58899
   162
  | str_of_mash_algorithm MaSh_kNN_Ext = "knn_ext"
blanchet@58899
   163
blanchet@58801
   164
fun scaled_avg [] = 0
blanchet@58801
   165
  | scaled_avg xs = Real.ceil (100000000.0 * fold (curry (op +)) xs 0.0) div length xs
blanchet@58801
   166
blanchet@58801
   167
fun avg [] = 0.0
blanchet@58801
   168
  | avg xs = fold (curry (op +)) xs 0.0 / Real.fromInt (length xs)
blanchet@58801
   169
blanchet@58801
   170
fun normalize_scores _ [] = []
blanchet@58801
   171
  | normalize_scores max_facts xs =
wneuper@59324
   172
    map (apsnd (curry (op *) (1.0 / avg (map snd (take max_facts xs))))) xs
blanchet@58801
   173
wneuper@59324
   174
fun mesh_facts maybe_distinct _ max_facts [(_, (sels, unks))] =
wneuper@59324
   175
    map fst (take max_facts sels) @ take (max_facts - length sels) unks
wneuper@59324
   176
    |> maybe_distinct
wneuper@59324
   177
  | mesh_facts _ fact_eq max_facts mess =
blanchet@58801
   178
    let
blanchet@58801
   179
      val mess = mess |> map (apsnd (apfst (normalize_scores max_facts)))
blanchet@58801
   180
blanchet@58801
   181
      fun score_in fact (global_weight, (sels, unks)) =
blanchet@58801
   182
        let val score_at = try (nth sels) #> Option.map (fn (_, score) => global_weight * score) in
blanchet@58801
   183
          (case find_index (curry fact_eq fact o fst) sels of
blanchet@58801
   184
            ~1 => if member fact_eq unks fact then NONE else SOME 0.0
blanchet@58801
   185
          | rank => score_at rank)
blanchet@58801
   186
        end
blanchet@58801
   187
blanchet@58801
   188
      fun weight_of fact = mess |> map_filter (score_in fact) |> scaled_avg
blanchet@58801
   189
    in
blanchet@58801
   190
      fold (union fact_eq o map fst o take max_facts o fst o snd) mess []
wneuper@59180
   191
      |> map (`weight_of) |> sort (int_ord o apply2 fst o swap)
blanchet@58801
   192
      |> map snd |> take max_facts
blanchet@58801
   193
    end
blanchet@58801
   194
blanchet@58801
   195
fun smooth_weight_of_fact rank = Math.pow (1.3, 15.5 - 0.2 * Real.fromInt rank) + 15.0 (* FUDGE *)
blanchet@58801
   196
fun steep_weight_of_fact rank = Math.pow (0.62, log2 (Real.fromInt (rank + 1))) (* FUDGE *)
blanchet@58801
   197
blanchet@58801
   198
fun weight_facts_smoothly facts = facts ~~ map smooth_weight_of_fact (0 upto length facts - 1)
blanchet@58801
   199
fun weight_facts_steeply facts = facts ~~ map steep_weight_of_fact (0 upto length facts - 1)
blanchet@58801
   200
blanchet@59002
   201
fun sort_array_suffix cmp needed a =
blanchet@58351
   202
  let
blanchet@58748
   203
    exception BOTTOM of int
blanchet@58748
   204
blanchet@58802
   205
    val al = Array.length a
blanchet@58802
   206
blanchet@58351
   207
    fun maxson l i =
blanchet@58359
   208
      let val i31 = i + i + i + 1 in
blanchet@58351
   209
        if i31 + 2 < l then
blanchet@58359
   210
          let val x = Unsynchronized.ref i31 in
wneuper@59451
   211
            if is_less (cmp (Array.sub (a, i31), Array.sub (a, i31 + 1))) then x := i31 + 1 else ();
wneuper@59451
   212
            if is_less (cmp (Array.sub (a, !x), Array.sub (a, i31 + 2))) then x := i31 + 2 else ();
blanchet@58351
   213
            !x
blanchet@58351
   214
          end
blanchet@58351
   215
        else
wneuper@59451
   216
          if i31 + 1 < l andalso is_less (cmp (Array.sub (a, i31), Array.sub (a, i31 + 1)))
blanchet@58351
   217
          then i31 + 1 else if i31 < l then i31 else raise BOTTOM i
blanchet@58351
   218
      end
blanchet@58351
   219
blanchet@58351
   220
    fun trickledown l i e =
blanchet@58371
   221
      let val j = maxson l i in
wneuper@59451
   222
        if is_greater (cmp (Array.sub (a, j), e)) then
blanchet@58371
   223
          (Array.update (a, i, Array.sub (a, j)); trickledown l j e)
blanchet@58371
   224
        else
blanchet@58371
   225
          Array.update (a, i, e)
blanchet@58351
   226
      end
blanchet@58351
   227
blanchet@58351
   228
    fun trickle l i e = trickledown l i e handle BOTTOM i => Array.update (a, i, e)
blanchet@58351
   229
blanchet@58351
   230
    fun bubbledown l i =
blanchet@58371
   231
      let val j = maxson l i in
blanchet@58371
   232
        Array.update (a, i, Array.sub (a, j));
blanchet@58351
   233
        bubbledown l j
blanchet@58351
   234
      end
blanchet@58351
   235
blanchet@58351
   236
    fun bubble l i = bubbledown l i handle BOTTOM i => i
blanchet@58351
   237
blanchet@58351
   238
    fun trickleup i e =
blanchet@58371
   239
      let val father = (i - 1) div 3 in
wneuper@59451
   240
        if is_less (cmp (Array.sub (a, father), e)) then
blanchet@58371
   241
          (Array.update (a, i, Array.sub (a, father));
blanchet@58371
   242
           if father > 0 then trickleup father e else Array.update (a, 0, e))
blanchet@58371
   243
        else
blanchet@58371
   244
          Array.update (a, i, e)
blanchet@58351
   245
      end
blanchet@58351
   246
blanchet@58697
   247
    fun for i = if i < 0 then () else (trickle al i (Array.sub (a, i)); for (i - 1))
blanchet@58351
   248
blanchet@58351
   249
    fun for2 i =
blanchet@59002
   250
      if i < Integer.max 2 (al - needed) then
blanchet@58371
   251
        ()
blanchet@58371
   252
      else
blanchet@58371
   253
        let val e = Array.sub (a, i) in
blanchet@58371
   254
          Array.update (a, i, Array.sub (a, 0));
blanchet@58371
   255
          trickleup (bubble i 0) e;
blanchet@58371
   256
          for2 (i - 1)
blanchet@58371
   257
        end
blanchet@58351
   258
  in
blanchet@58697
   259
    for (((al + 1) div 3) - 1);
blanchet@58697
   260
    for2 (al - 1);
blanchet@58697
   261
    if al > 1 then
blanchet@58371
   262
      let val e = Array.sub (a, 1) in
blanchet@58371
   263
        Array.update (a, 1, Array.sub (a, 0));
blanchet@58351
   264
        Array.update (a, 0, e)
blanchet@58351
   265
      end
blanchet@58371
   266
    else
blanchet@58371
   267
      ()
blanchet@58351
   268
  end
blanchet@58351
   269
blanchet@59002
   270
fun rev_sort_list_prefix cmp needed xs =
blanchet@58802
   271
  let val ary = Array.fromList xs in
blanchet@59002
   272
    sort_array_suffix cmp needed ary;
blanchet@59002
   273
    Array.foldl (op ::) [] ary
blanchet@58802
   274
  end
blanchet@58802
   275
blanchet@58802
   276
wneuper@59324
   277
(*** Convenience functions for synchronized access ***)
wneuper@59324
   278
wneuper@59324
   279
fun synchronized_timed_value var time_limit =
wneuper@59324
   280
  Synchronized.timed_access var time_limit (fn value => SOME (value, value))
wneuper@59324
   281
fun synchronized_timed_change_result var time_limit f =
wneuper@59324
   282
  Synchronized.timed_access var time_limit (SOME o f)
wneuper@59324
   283
fun synchronized_timed_change var time_limit f =
wneuper@59324
   284
  synchronized_timed_change_result var time_limit (fn x => ((), f x))
wneuper@59324
   285
wneuper@59324
   286
fun mash_time_limit _ = SOME (seconds 0.1)
wneuper@59324
   287
wneuper@59324
   288
blanchet@58802
   289
(*** Isabelle-agnostic machine learning ***)
blanchet@58802
   290
blanchet@58802
   291
structure MaSh =
blanchet@58802
   292
struct
blanchet@58802
   293
walther@59606
   294
fun select_fact_idxs (big_number : real) recommends =
blanchet@58706
   295
  List.app (fn at =>
blanchet@58706
   296
    let val (j, ov) = Array.sub (recommends, at) in
blanchet@58713
   297
      Array.update (recommends, at, (j, big_number + ov))
blanchet@58706
   298
    end)
blanchet@58706
   299
blanchet@58716
   300
fun wider_array_of_vector init vec =
blanchet@58716
   301
  let val ary = Array.array init in
blanchet@58716
   302
    Array.copyVec {src = vec, dst = ary, di = 0};
blanchet@58716
   303
    ary
blanchet@58716
   304
  end
blanchet@58716
   305
blanchet@58873
   306
val nb_def_prior_weight = 1000 (* FUDGE *)
blanchet@58437
   307
blanchet@58716
   308
fun learn_facts (tfreq0, sfreq0, dffreq0) num_facts0 num_facts num_feats depss featss =
blanchet@58696
   309
  let
blanchet@58716
   310
    val tfreq = wider_array_of_vector (num_facts, 0) tfreq0
blanchet@58716
   311
    val sfreq = wider_array_of_vector (num_facts, Inttab.empty) sfreq0
blanchet@58716
   312
    val dffreq = wider_array_of_vector (num_feats, 0) dffreq0
blanchet@58716
   313
blanchet@58716
   314
    fun learn_one th feats deps =
blanchet@58696
   315
      let
blanchet@58697
   316
        fun add_th weight t =
blanchet@58697
   317
          let
blanchet@58697
   318
            val im = Array.sub (sfreq, t)
blanchet@58716
   319
            fun fold_fn s = Inttab.map_default (s, 0) (Integer.add weight)
blanchet@58697
   320
          in
blanchet@58710
   321
            map_array_at tfreq (Integer.add weight) t;
blanchet@58697
   322
            Array.update (sfreq, t, fold fold_fn feats im)
blanchet@58697
   323
          end
blanchet@58697
   324
blanchet@58710
   325
        val add_sym = map_array_at dffreq (Integer.add 1)
blanchet@58696
   326
      in
blanchet@58697
   327
        add_th nb_def_prior_weight th;
blanchet@58697
   328
        List.app (add_th 1) deps;
blanchet@58697
   329
        List.app add_sym feats
blanchet@58696
   330
      end
blanchet@58696
   331
blanchet@58444
   332
    fun for i =
blanchet@58725
   333
      if i = num_facts then ()
blanchet@58725
   334
      else (learn_one i (Vector.sub (featss, i)) (Vector.sub (depss, i)); for (i + 1))
blanchet@58444
   335
  in
blanchet@58725
   336
    for num_facts0;
blanchet@58716
   337
    (Array.vector tfreq, Array.vector sfreq, Array.vector dffreq)
blanchet@58444
   338
  end
blanchet@58371
   339
walther@59606
   340
fun naive_bayes (tfreq, sfreq, dffreq) num_facts max_suggs fact_idxs goal_feats =
blanchet@58620
   341
  let
blanchet@58873
   342
    val tau = 0.2 (* FUDGE *)
blanchet@58873
   343
    val pos_weight = 5.0 (* FUDGE *)
blanchet@58873
   344
    val def_val = ~18.0 (* FUDGE *)
blanchet@58873
   345
    val init_val = 30.0 (* FUDGE *)
blanchet@58466
   346
blanchet@58708
   347
    val ln_afreq = Math.ln (Real.fromInt num_facts)
blanchet@58716
   348
    val idf = Vector.map (fn i => ln_afreq - Math.ln (Real.fromInt i)) dffreq
blanchet@58708
   349
blanchet@58701
   350
    fun tfidf feat = Vector.sub (idf, feat)
blanchet@58444
   351
blanchet@58444
   352
    fun log_posterior i =
blanchet@58371
   353
      let
blanchet@58716
   354
        val tfreq = Real.fromInt (Vector.sub (tfreq, i))
blanchet@58439
   355
blanchet@58873
   356
        fun add_feat (f, fw0) (res, sfh) =
blanchet@58444
   357
          (case Inttab.lookup sfh f of
blanchet@58444
   358
            SOME sf =>
blanchet@58743
   359
            (res + fw0 * tfidf f * Math.ln (pos_weight * Real.fromInt sf / tfreq),
blanchet@58444
   360
             Inttab.delete f sfh)
blanchet@58751
   361
          | NONE => (res + fw0 * tfidf f * def_val, sfh))
blanchet@58439
   362
blanchet@58873
   363
        val (res, sfh) = fold add_feat goal_feats (init_val * Math.ln tfreq, Vector.sub (sfreq, i))
blanchet@58439
   364
blanchet@58751
   365
        fun fold_sfh (f, sf) sow =
wneuper@59180
   366
          sow + tfidf f * Math.ln (1.0 - Real.fromInt (sf - 1) / tfreq)
blanchet@58439
   367
blanchet@58444
   368
        val sum_of_weights = Inttab.fold fold_sfh sfh 0.0
blanchet@58371
   369
      in
blanchet@58620
   370
        res + tau * sum_of_weights
blanchet@58439
   371
      end
blanchet@58371
   372
blanchet@58698
   373
    val posterior = Array.tabulate (num_facts, (fn j => (j, log_posterior j)))
blanchet@58444
   374
blanchet@58706
   375
    fun ret at acc =
blanchet@58706
   376
      if at = num_facts then acc else ret (at + 1) (Array.sub (posterior, at) :: acc)
blanchet@58371
   377
  in
walther@59606
   378
    select_fact_idxs 100000.0 posterior fact_idxs;
wneuper@59180
   379
    sort_array_suffix (Real.compare o apply2 snd) max_suggs posterior;
blanchet@58706
   380
    ret (Integer.max 0 (num_facts - max_suggs)) []
blanchet@58371
   381
  end
blanchet@58371
   382
wneuper@59180
   383
val initial_k = 0
blanchet@58800
   384
walther@59606
   385
fun k_nearest_neighbors dffreq num_facts num_feats depss featss max_suggs fact_idxs goal_feats =
blanchet@58800
   386
  let
blanchet@58800
   387
    exception EXIT of unit
blanchet@58800
   388
blanchet@58800
   389
    val ln_afreq = Math.ln (Real.fromInt num_facts)
blanchet@58800
   390
    fun tfidf feat = ln_afreq - Math.ln (Real.fromInt (Vector.sub (dffreq, feat)))
blanchet@58800
   391
blanchet@58800
   392
    val overlaps_sqr = Array.tabulate (num_facts, rpair 0.0)
blanchet@58800
   393
blanchet@58801
   394
    val feat_facts = Array.array (num_feats, [])
blanchet@58801
   395
    val _ = Vector.foldl (fn (feats, fact) =>
blanchet@58801
   396
      (List.app (map_array_at feat_facts (cons fact)) feats; fact + 1)) 0 featss
blanchet@58801
   397
blanchet@58800
   398
    fun do_feat (s, sw0) =
blanchet@58800
   399
      let
blanchet@58800
   400
        val sw = sw0 * tfidf s
wneuper@59180
   401
        val w6 = Math.pow (sw, 6.0 (* FUDGE *))
blanchet@58800
   402
blanchet@58800
   403
        fun inc_overlap j =
blanchet@58800
   404
          let val (_, ov) = Array.sub (overlaps_sqr, j) in
blanchet@58873
   405
            Array.update (overlaps_sqr, j, (j, w6 + ov))
blanchet@58800
   406
          end
blanchet@58800
   407
      in
blanchet@58800
   408
        List.app inc_overlap (Array.sub (feat_facts, s))
blanchet@58800
   409
      end
blanchet@58800
   410
blanchet@58800
   411
    val _ = List.app do_feat goal_feats
wneuper@59180
   412
    val _ = sort_array_suffix (Real.compare o apply2 snd) num_facts overlaps_sqr
blanchet@58800
   413
    val no_recommends = Unsynchronized.ref 0
blanchet@58800
   414
    val recommends = Array.tabulate (num_facts, rpair 0.0)
blanchet@58800
   415
    val age = Unsynchronized.ref 500000000.0
blanchet@58800
   416
blanchet@58800
   417
    fun inc_recommend v j =
blanchet@58800
   418
      let val (_, ov) = Array.sub (recommends, j) in
blanchet@58800
   419
        if ov <= 0.0 then
blanchet@58800
   420
          (no_recommends := !no_recommends + 1; Array.update (recommends, j, (j, !age + ov)))
wneuper@59180
   421
        else
wneuper@59180
   422
          Array.update (recommends, j, (j, v + ov))
blanchet@58800
   423
      end
blanchet@58800
   424
blanchet@58800
   425
    val k = Unsynchronized.ref 0
blanchet@58800
   426
    fun do_k k =
blanchet@58800
   427
      if k >= num_facts then
blanchet@58800
   428
        raise EXIT ()
blanchet@58800
   429
      else
blanchet@58800
   430
        let
blanchet@58873
   431
          val deps_factor = 2.7 (* FUDGE *)
blanchet@58800
   432
          val (j, o2) = Array.sub (overlaps_sqr, num_facts - k - 1)
blanchet@58873
   433
          val _ = inc_recommend o2 j
blanchet@58800
   434
          val ds = Vector.sub (depss, j)
blanchet@58800
   435
          val l = Real.fromInt (length ds)
blanchet@58800
   436
        in
blanchet@58873
   437
          List.app (inc_recommend (deps_factor * o2 / l)) ds
blanchet@58800
   438
        end
blanchet@58800
   439
blanchet@58800
   440
    fun while1 () =
wneuper@59180
   441
      if !k = initial_k + 1 then () else (do_k (!k); k := !k + 1; while1 ())
blanchet@58800
   442
      handle EXIT () => ()
blanchet@58800
   443
blanchet@58800
   444
    fun while2 () =
blanchet@58800
   445
      if !no_recommends >= max_suggs then ()
blanchet@58800
   446
      else (do_k (!k); k := !k + 1; age := !age - 10000.0; while2 ())
blanchet@58800
   447
      handle EXIT () => ()
blanchet@58800
   448
blanchet@58800
   449
    fun ret acc at =
blanchet@58800
   450
      if at = num_facts then acc else ret (Array.sub (recommends, at) :: acc) (at + 1)
blanchet@58800
   451
  in
blanchet@58800
   452
    while1 ();
blanchet@58800
   453
    while2 ();
walther@59606
   454
    select_fact_idxs 1000000000.0 recommends fact_idxs;
wneuper@59180
   455
    sort_array_suffix (Real.compare o apply2 snd) max_suggs recommends;
blanchet@58800
   456
    ret [] (Integer.max 0 (num_facts - max_suggs))
blanchet@58800
   457
  end
blanchet@58800
   458
blanchet@58467
   459
(* experimental *)
blanchet@58718
   460
fun external_tool tool max_suggs learns goal_feats =
blanchet@58633
   461
  let
blanchet@58639
   462
    val ser = string_of_int (serial ()) (* poor person's attempt at thread-safety *)
blanchet@58639
   463
    val ocs = TextIO.openOut ("adv_syms" ^ ser)
blanchet@58639
   464
    val ocd = TextIO.openOut ("adv_deps" ^ ser)
blanchet@58639
   465
    val ocq = TextIO.openOut ("adv_seq" ^ ser)
blanchet@58639
   466
    val occ = TextIO.openOut ("adv_conj" ^ ser)
blanchet@58638
   467
blanchet@58633
   468
    fun os oc s = TextIO.output (oc, s)
blanchet@58638
   469
blanchet@58639
   470
    fun ol _ _ _ [] = ()
blanchet@58639
   471
      | ol _ f _ [e] = f e
blanchet@58633
   472
      | ol oc f sep (h :: t) = (f h; os oc sep; ol oc f sep t)
blanchet@58638
   473
blanchet@58636
   474
    fun do_learn (name, feats, deps) =
blanchet@58699
   475
      (os ocs name; os ocs ":"; ol ocs (os ocs o quote) ", " feats; os ocs "\n";
blanchet@58639
   476
       os ocd name; os ocd ":"; ol ocd (os ocd) " " deps; os ocd "\n"; os ocq name; os ocq "\n")
blanchet@58638
   477
blanchet@58633
   478
    fun forkexec no =
blanchet@58633
   479
      let
blanchet@58633
   480
        val cmd =
blanchet@58639
   481
          "~/misc/" ^ tool ^ " adv_syms" ^ ser ^ " adv_deps" ^ ser ^ " " ^ string_of_int no ^
blanchet@58639
   482
          " adv_seq" ^ ser ^ " < adv_conj" ^ ser
blanchet@58633
   483
      in
blanchet@58633
   484
        fst (Isabelle_System.bash_output cmd)
blanchet@58633
   485
        |> space_explode " "
blanchet@58636
   486
        |> filter_out (curry (op =) "")
blanchet@58633
   487
      end
blanchet@58633
   488
  in
blanchet@58743
   489
    (List.app do_learn learns; ol occ (os occ o quote) ", " (map fst goal_feats);
blanchet@58639
   490
     TextIO.closeOut ocs; TextIO.closeOut ocd; TextIO.closeOut ocq; TextIO.closeOut occ;
blanchet@58636
   491
     forkexec max_suggs)
blanchet@58633
   492
  end
blanchet@58633
   493
blanchet@58903
   494
fun k_nearest_neighbors_ext max_suggs =
wneuper@59180
   495
  external_tool ("newknn/knn" ^ " " ^ string_of_int initial_k) max_suggs
blanchet@58903
   496
fun naive_bayes_ext max_suggs = external_tool "predict/nbayes" max_suggs
blanchet@58639
   497
blanchet@58874
   498
fun query_external ctxt algorithm max_suggs learns goal_feats =
blanchet@58774
   499
  (trace_msg ctxt (fn () => "MaSh query external " ^ commas (map fst goal_feats));
blanchet@58874
   500
   (case algorithm of
blanchet@58800
   501
     MaSh_NB_Ext => naive_bayes_ext max_suggs learns goal_feats
blanchet@58800
   502
   | MaSh_kNN_Ext => k_nearest_neighbors_ext max_suggs learns goal_feats))
blanchet@58716
   503
blanchet@58874
   504
fun query_internal ctxt algorithm num_facts num_feats (fact_names, featss, depss)
walther@59606
   505
    (freqs as (_, _, dffreq)) fact_idxs max_suggs goal_feats int_goal_feats =
blanchet@58801
   506
  let
blanchet@58801
   507
    fun nb () =
walther@59606
   508
      naive_bayes freqs num_facts max_suggs fact_idxs int_goal_feats
blanchet@58801
   509
      |> map fst
blanchet@58801
   510
    fun knn () =
walther@59606
   511
      k_nearest_neighbors dffreq num_facts num_feats depss featss max_suggs fact_idxs int_goal_feats
blanchet@58801
   512
      |> map fst
blanchet@58801
   513
  in
blanchet@58801
   514
    (trace_msg ctxt (fn () => "MaSh query internal " ^ commas (map fst goal_feats) ^ " from {" ^
blanchet@58801
   515
       elide_string 1000 (space_implode " " (Vector.foldr (op ::) [] fact_names)) ^ "}");
blanchet@58874
   516
     (case algorithm of
blanchet@58801
   517
       MaSh_NB => nb ()
blanchet@58801
   518
     | MaSh_kNN => knn ()
blanchet@58801
   519
     | MaSh_NB_kNN =>
wneuper@59324
   520
       mesh_facts I (op =) max_suggs
blanchet@58894
   521
         [(0.5 (* FUDGE *), (weight_facts_steeply (nb ()), [])),
blanchet@58894
   522
          (0.5 (* FUDGE *), (weight_facts_steeply (knn ()), []))])
blanchet@58801
   523
     |> map (curry Vector.sub fact_names))
blanchet@58801
   524
   end
blanchet@58349
   525
blanchet@58349
   526
end;
blanchet@58349
   527
blanchet@58349
   528
blanchet@58774
   529
(*** Persistent, stringly-typed state ***)
blanchet@58774
   530
blanchet@58774
   531
fun meta_char c =
blanchet@58774
   532
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse c = #")" orelse
wneuper@59324
   533
     c = #"," orelse c = #"'" then
blanchet@58774
   534
    String.str c
blanchet@58774
   535
  else
blanchet@58774
   536
    (* fixed width, in case more digits follow *)
blanchet@58774
   537
    "%" ^ stringN_of_int 3 (Char.ord c)
blanchet@58774
   538
blanchet@58774
   539
fun unmeta_chars accum [] = String.implode (rev accum)
blanchet@58774
   540
  | unmeta_chars accum (#"%" :: d1 :: d2 :: d3 :: cs) =
blanchet@58774
   541
    (case Int.fromString (String.implode [d1, d2, d3]) of
blanchet@58774
   542
      SOME n => unmeta_chars (Char.chr n :: accum) cs
blanchet@58774
   543
    | NONE => "" (* error *))
blanchet@58774
   544
  | unmeta_chars _ (#"%" :: _) = "" (* error *)
blanchet@58774
   545
  | unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs
blanchet@58774
   546
blanchet@58774
   547
val encode_str = String.translate meta_char
wneuper@59324
   548
val encode_strs = map encode_str #> space_implode " "
blanchet@58774
   549
wneuper@59324
   550
fun decode_str s =
wneuper@59324
   551
  if String.isSubstring "%" s then unmeta_chars [] (String.explode s) else s;
wneuper@59324
   552
wneuper@59324
   553
fun decode_strs s =
wneuper@59324
   554
  space_explode " " s |> String.isSubstring "%" s ? map decode_str;
blanchet@51326
   555
blanchet@58347
   556
datatype proof_kind = Isar_Proof | Automatic_Proof | Isar_Proof_wegen_Prover_Flop
blanchet@51326
   557
blanchet@51326
   558
fun str_of_proof_kind Isar_Proof = "i"
blanchet@51499
   559
  | str_of_proof_kind Automatic_Proof = "a"
blanchet@51499
   560
  | str_of_proof_kind Isar_Proof_wegen_Prover_Flop = "x"
blanchet@51326
   561
blanchet@58347
   562
fun proof_kind_of_str "a" = Automatic_Proof
blanchet@51499
   563
  | proof_kind_of_str "x" = Isar_Proof_wegen_Prover_Flop
blanchet@58347
   564
  | proof_kind_of_str _ (* "i" *) = Isar_Proof
blanchet@51326
   565
blanchet@58353
   566
fun add_edge_to name parent =
blanchet@58353
   567
  Graph.default_node (parent, (Isar_Proof, [], []))
blanchet@58353
   568
  #> Graph.add_edge (parent, name)
blanchet@58353
   569
blanchet@59003
   570
fun add_node kind name parents feats deps (accum as (access_G, (fact_xtab, feat_xtab), learns)) =
blanchet@59003
   571
  let val fact_xtab' = add_to_xtab name fact_xtab in
blanchet@59003
   572
    ((Graph.new_node (name, (kind, feats, deps)) access_G
blanchet@59003
   573
      handle Graph.DUP _ => Graph.map_node name (K (kind, feats, deps)) access_G)
blanchet@59003
   574
     |> fold (add_edge_to name) parents,
blanchet@59003
   575
     (fact_xtab', fold maybe_add_to_xtab feats feat_xtab),
blanchet@59003
   576
     (name, feats, deps) :: learns)
blanchet@59003
   577
  end
blanchet@59003
   578
  handle Symtab.DUP _ => accum (* robustness (in case the state file violates the invariant) *)
blanchet@58353
   579
blanchet@51326
   580
fun try_graph ctxt when def f =
blanchet@51326
   581
  f ()
blanchet@58347
   582
  handle
blanchet@58347
   583
    Graph.CYCLES (cycle :: _) =>
blanchet@58347
   584
    (trace_msg ctxt (fn () => "Cycle involving " ^ commas cycle ^ " when " ^ when); def)
blanchet@58347
   585
  | Graph.DUP name =>
blanchet@58347
   586
    (trace_msg ctxt (fn () => "Duplicate fact " ^ quote name ^ " when " ^ when); def)
blanchet@58347
   587
  | Graph.UNDEF name =>
blanchet@58347
   588
    (trace_msg ctxt (fn () => "Unknown fact " ^ quote name ^ " when " ^ when); def)
blanchet@58347
   589
  | exn =>
blanchet@58347
   590
    if Exn.is_interrupt exn then
wneuper@59324
   591
      Exn.reraise exn
blanchet@58347
   592
    else
blanchet@58347
   593
      (trace_msg ctxt (fn () => "Internal error when " ^ when ^ ":\n" ^ Runtime.exn_message exn);
blanchet@58347
   594
       def)
blanchet@51326
   595
blanchet@51326
   596
fun graph_info G =
blanchet@51326
   597
  string_of_int (length (Graph.keys G)) ^ " node(s), " ^
blanchet@58348
   598
  string_of_int (fold (Integer.add o length o snd) (Graph.dest G) 0) ^ " edge(s), " ^
blanchet@51326
   599
  string_of_int (length (Graph.maximals G)) ^ " maximal"
blanchet@51326
   600
blanchet@58887
   601
type ffds = string vector * int list vector * int list vector
blanchet@58887
   602
type freqs = int vector * int Inttab.table vector * int vector
blanchet@58887
   603
blanchet@54232
   604
type mash_state =
blanchet@58700
   605
  {access_G : (proof_kind * string list * string list) Graph.T,
blanchet@58716
   606
   xtabs : xtab * xtab,
blanchet@58887
   607
   ffds : ffds,
blanchet@58887
   608
   freqs : freqs,
blanchet@58707
   609
   dirty_facts : string list option}
blanchet@51326
   610
blanchet@58720
   611
val empty_xtabs = (empty_xtab, empty_xtab)
blanchet@58887
   612
val empty_ffds = (Vector.fromList [], Vector.fromList [], Vector.fromList []) : ffds
blanchet@58887
   613
val empty_freqs = (Vector.fromList [], Vector.fromList [], Vector.fromList []) : freqs
blanchet@58720
   614
blanchet@58713
   615
val empty_state =
blanchet@58713
   616
  {access_G = Graph.empty,
blanchet@58720
   617
   xtabs = empty_xtabs,
blanchet@58720
   618
   ffds = empty_ffds,
blanchet@58720
   619
   freqs = empty_freqs,
blanchet@58713
   620
   dirty_facts = SOME []} : mash_state
blanchet@58713
   621
blanchet@58887
   622
fun recompute_ffds_freqs_from_learns (learns : (string * string list * string list) list)
blanchet@58887
   623
    ((num_facts, fact_tab), (num_feats, feat_tab)) num_facts0 (fact_names0, featss0, depss0) freqs0 =
blanchet@58721
   624
  let
blanchet@58722
   625
    val fact_names = Vector.concat [fact_names0, Vector.fromList (map #1 learns)]
blanchet@58722
   626
    val featss = Vector.concat [featss0,
blanchet@58722
   627
      Vector.fromList (map (map_filter (Symtab.lookup feat_tab) o #2) learns)]
blanchet@58722
   628
    val depss = Vector.concat [depss0,
blanchet@58722
   629
      Vector.fromList (map (map_filter (Symtab.lookup fact_tab) o #3) learns)]
blanchet@58721
   630
  in
blanchet@58721
   631
    ((fact_names, featss, depss),
blanchet@58773
   632
     MaSh.learn_facts freqs0 num_facts0 num_facts num_feats depss featss)
blanchet@58721
   633
  end
blanchet@58721
   634
blanchet@58720
   635
fun reorder_learns (num_facts, fact_tab) learns =
blanchet@58720
   636
  let val ary = Array.array (num_facts, ("", [], [])) in
blanchet@58720
   637
    List.app (fn learn as (fact, _, _) =>
blanchet@58720
   638
        Array.update (ary, the (Symtab.lookup fact_tab fact), learn))
blanchet@58720
   639
      learns;
blanchet@58720
   640
    Array.foldr (op ::) [] ary
blanchet@58720
   641
  end
blanchet@58720
   642
blanchet@58722
   643
fun recompute_ffds_freqs_from_access_G access_G (xtabs as (fact_xtab, _)) =
blanchet@58720
   644
  let
blanchet@58720
   645
    val learns =
blanchet@58720
   646
      Graph.schedule (fn _ => fn (fact, (_, feats, deps)) => (fact, feats, deps)) access_G
blanchet@58720
   647
      |> reorder_learns fact_xtab
blanchet@58720
   648
  in
blanchet@58722
   649
    recompute_ffds_freqs_from_learns learns xtabs 0 empty_ffds empty_freqs
blanchet@58720
   650
  end
blanchet@51326
   651
blanchet@51326
   652
local
blanchet@51326
   653
walther@59606
   654
val version = "*** MaSh version 20190121 ***"
blanchet@51372
   655
blanchet@54232
   656
exception FILE_VERSION_TOO_NEW of unit
blanchet@51326
   657
blanchet@51326
   658
fun extract_node line =
blanchet@56628
   659
  (case space_explode ":" line of
blanchet@58347
   660
    [head, tail] =>
blanchet@58347
   661
    (case (space_explode " " head, map (unprefix " ") (space_explode ";" tail)) of
blanchet@58347
   662
      ([kind, name], [parents, feats, deps]) =>
blanchet@58700
   663
      SOME (proof_kind_of_str kind, decode_str name, decode_strs parents, decode_strs feats,
blanchet@58352
   664
        decode_strs deps)
blanchet@56628
   665
    | _ => NONE)
blanchet@56628
   666
  | _ => NONE)
blanchet@51326
   667
wneuper@59324
   668
fun would_load_state (memory_time, _) =
wneuper@59324
   669
  let val path = state_file () in
wneuper@59324
   670
    (case try OS.FileSys.modTime (File.platform_path path) of
wneuper@59324
   671
      NONE => false
wneuper@59324
   672
    | SOME disk_time => memory_time < disk_time)
wneuper@59324
   673
  end;
wneuper@59324
   674
blanchet@58773
   675
fun load_state ctxt (time_state as (memory_time, _)) =
blanchet@58894
   676
  let val path = state_file () in
wneuper@59324
   677
    (case try OS.FileSys.modTime (File.platform_path path) of
blanchet@58418
   678
      NONE => time_state
blanchet@58418
   679
    | SOME disk_time =>
wneuper@59324
   680
      if memory_time >= disk_time then
blanchet@58418
   681
        time_state
blanchet@58418
   682
      else
blanchet@58418
   683
        (disk_time,
blanchet@58418
   684
         (case try File.read_lines path of
blanchet@58418
   685
           SOME (version' :: node_lines) =>
blanchet@58418
   686
           let
blanchet@58418
   687
             fun extract_line_and_add_node line =
blanchet@58418
   688
               (case extract_node line of
blanchet@58418
   689
                 NONE => I (* should not happen *)
blanchet@58418
   690
               | SOME (kind, name, parents, feats, deps) => add_node kind name parents feats deps)
blanchet@58353
   691
blanchet@58721
   692
             val empty_G_etc = (Graph.empty, empty_xtabs, [])
blanchet@58721
   693
blanchet@58721
   694
             val (access_G, xtabs, rev_learns) =
blanchet@58418
   695
               (case string_ord (version', version) of
blanchet@58418
   696
                 EQUAL =>
blanchet@58721
   697
                 try_graph ctxt "loading state" empty_G_etc
blanchet@58721
   698
                   (fn () => fold extract_line_and_add_node node_lines empty_G_etc)
blanchet@58894
   699
               | LESS => (remove_state_file (); empty_G_etc) (* cannot parse old file *)
blanchet@58418
   700
               | GREATER => raise FILE_VERSION_TOO_NEW ())
blanchet@58720
   701
blanchet@58722
   702
             val (ffds, freqs) =
blanchet@58722
   703
               recompute_ffds_freqs_from_learns (rev rev_learns) xtabs 0 empty_ffds empty_freqs
blanchet@58418
   704
           in
blanchet@58418
   705
             trace_msg ctxt (fn () => "Loaded fact graph (" ^ graph_info access_G ^ ")");
blanchet@58720
   706
             {access_G = access_G, xtabs = xtabs, ffds = ffds, freqs = freqs, dirty_facts = SOME []}
blanchet@58418
   707
           end
blanchet@58418
   708
         | _ => empty_state)))
blanchet@58418
   709
  end
blanchet@51326
   710
blanchet@58347
   711
fun str_of_entry (kind, name, parents, feats, deps) =
blanchet@58347
   712
  str_of_proof_kind kind ^ " " ^ encode_str name ^ ": " ^ encode_strs parents ^ "; " ^
blanchet@58700
   713
  encode_strs feats ^ "; " ^ encode_strs deps ^ "\n"
blanchet@58347
   714
blanchet@58707
   715
fun save_state _ (time_state as (_, {dirty_facts = SOME [], ...})) = time_state
blanchet@58720
   716
  | save_state ctxt (memory_time, {access_G, xtabs, ffds, freqs, dirty_facts}) =
blanchet@51326
   717
    let
blanchet@58347
   718
      fun append_entry (name, ((kind, feats, deps), (parents, _))) =
blanchet@58347
   719
        cons (kind, name, Graph.Keys.dest parents, feats, deps)
blanchet@58347
   720
blanchet@58894
   721
      val path = state_file ()
blanchet@58707
   722
      val dirty_facts' =
wneuper@59324
   723
        (case try OS.FileSys.modTime (File.platform_path path) of
blanchet@58418
   724
          NONE => NONE
wneuper@59324
   725
        | SOME disk_time => if disk_time <= memory_time then dirty_facts else NONE)
blanchet@51326
   726
      val (banner, entries) =
blanchet@58707
   727
        (case dirty_facts' of
blanchet@56628
   728
          SOME names => (NONE, fold (append_entry o Graph.get_entry access_G) names [])
blanchet@56628
   729
        | NONE => (SOME (version ^ "\n"), Graph.fold append_entry access_G []))
blanchet@51326
   730
    in
blanchet@58773
   731
      (case banner of SOME s => File.write path s | NONE => ();
blanchet@58773
   732
       entries |> chunk_list 500 |> List.app (File.append path o implode o map str_of_entry))
blanchet@58773
   733
      handle IO.Io _ => ();
blanchet@51326
   734
      trace_msg ctxt (fn () =>
blanchet@58347
   735
        "Saved fact graph (" ^ graph_info access_G ^
blanchet@58707
   736
        (case dirty_facts of
blanchet@58707
   737
          SOME dirty_facts => "; " ^ string_of_int (length dirty_facts) ^ " dirty fact(s)"
blanchet@58347
   738
        | _ => "") ^  ")");
blanchet@58720
   739
      (Time.now (),
blanchet@58720
   740
       {access_G = access_G, xtabs = xtabs, ffds = ffds, freqs = freqs, dirty_facts = SOME []})
blanchet@51326
   741
    end
blanchet@51326
   742
blanchet@58707
   743
val global_state = Synchronized.var "Sledgehammer_MaSh.global_state" (Time.zeroTime, empty_state)
blanchet@51326
   744
blanchet@51326
   745
in
blanchet@51326
   746
blanchet@58773
   747
fun map_state ctxt f =
wneuper@59324
   748
  (trace_msg ctxt (fn () => "Changing MaSh state");
wneuper@59324
   749
   synchronized_timed_change global_state mash_time_limit
wneuper@59324
   750
     (load_state ctxt ##> f #> save_state ctxt))
wneuper@59324
   751
  |> ignore
blanchet@54232
   752
  handle FILE_VERSION_TOO_NEW () => ()
blanchet@51326
   753
blanchet@58773
   754
fun peek_state ctxt =
wneuper@59324
   755
  (trace_msg ctxt (fn () => "Peeking at MaSh state");
wneuper@59324
   756
   (case synchronized_timed_value global_state mash_time_limit of
wneuper@59324
   757
     NONE => NONE
wneuper@59324
   758
   | SOME state => if would_load_state state then NONE else SOME state))
blanchet@51326
   759
wneuper@59324
   760
fun get_state ctxt =
wneuper@59324
   761
  (trace_msg ctxt (fn () => "Retrieving MaSh state");
wneuper@59324
   762
   synchronized_timed_change_result global_state mash_time_limit
wneuper@59324
   763
     (perhaps (try (load_state ctxt)) #> `snd))
wneuper@59324
   764
wneuper@59324
   765
fun clear_state ctxt =
wneuper@59324
   766
  (trace_msg ctxt (fn () => "Clearing MaSh state");
wneuper@59324
   767
   Synchronized.change global_state (fn _ => (remove_state_file (); (Time.zeroTime, empty_state))))
blanchet@51326
   768
blanchet@51326
   769
end
blanchet@51326
   770
blanchet@51326
   771
blanchet@51326
   772
(*** Isabelle helpers ***)
blanchet@51326
   773
wneuper@59324
   774
fun crude_printed_term size t =
wneuper@59324
   775
  let
wneuper@59324
   776
    fun term _ (res, 0) = (res, 0)
wneuper@59324
   777
      | term (t $ u) (res, size) =
wneuper@59324
   778
        let
wneuper@59324
   779
          val (res, size) = term t (res ^ "(", size)
wneuper@59324
   780
          val (res, size) = term u (res ^ " ", size)
wneuper@59324
   781
        in
wneuper@59324
   782
          (res ^ ")", size)
wneuper@59324
   783
        end
wneuper@59324
   784
      | term (Abs (s, _, t)) (res, size) = term t (res ^ "%" ^ s ^ ".", size - 1)
wneuper@59324
   785
      | term (Bound n) (res, size) = (res ^ "#" ^ string_of_int n, size - 1)
wneuper@59324
   786
      | term (Const (s, _)) (res, size) = (res ^ Long_Name.base_name s, size - 1)
wneuper@59324
   787
      | term (Free (s, _)) (res, size) = (res ^ s, size - 1)
wneuper@59324
   788
      | term (Var ((s, _), _)) (res, size) = (res ^ s, size - 1)
wneuper@59324
   789
  in
wneuper@59324
   790
    fst (term t ("", size))
wneuper@59324
   791
  end
blanchet@52318
   792
blanchet@51639
   793
fun nickname_of_thm th =
blanchet@49409
   794
  if Thm.has_name_hint th then
blanchet@49409
   795
    let val hint = Thm.get_name_hint th in
blanchet@51737
   796
      (* There must be a better way to detect local facts. *)
wneuper@59180
   797
      (case Long_Name.dest_local hint of
wneuper@59324
   798
        SOME suf =>
wneuper@59324
   799
        Long_Name.implode [Thm.theory_name th, suf, crude_printed_term 25 (Thm.prop_of th)]
blanchet@56628
   800
      | NONE => hint)
blanchet@49409
   801
    end
blanchet@49409
   802
  else
wneuper@59324
   803
    crude_printed_term 50 (Thm.prop_of th)
blanchet@49393
   804
blanchet@52271
   805
fun find_suggested_facts ctxt facts =
blanchet@49345
   806
  let
blanchet@52271
   807
    fun add (fact as (_, th)) = Symtab.default (nickname_of_thm th, fact)
blanchet@52271
   808
    val tab = fold add facts Symtab.empty
blanchet@52271
   809
    fun lookup nick =
blanchet@52271
   810
      Symtab.lookup tab nick
blanchet@58348
   811
      |> tap (fn NONE => trace_msg ctxt (fn () => "Cannot find " ^ quote nick) | _ => ())
blanchet@52271
   812
  in map_filter lookup end
blanchet@49326
   813
blanchet@58746
   814
fun free_feature_of s = "f" ^ s
blanchet@58746
   815
fun thy_feature_of s = "y" ^ s
blanchet@58746
   816
fun type_feature_of s = "t" ^ s
blanchet@58746
   817
fun class_feature_of s = "s" ^ s
blanchet@58746
   818
val local_feature = "local"
blanchet@49266
   819
wneuper@59324
   820
fun crude_thm_ord ctxt =
wneuper@59324
   821
  let
wneuper@59324
   822
    val ancestor_lengths =
wneuper@59324
   823
      fold (fn thy => Symtab.update (Context.theory_name thy, length (Context.ancestors_of thy)))
wneuper@59324
   824
        (Theory.nodes_of (Proof_Context.theory_of ctxt)) Symtab.empty
wneuper@59324
   825
    val ancestor_length = Symtab.lookup ancestor_lengths o Context.theory_id_name
blanchet@49339
   826
wneuper@59324
   827
    fun crude_theory_ord p =
wneuper@59324
   828
      if Context.eq_thy_id p then EQUAL
wneuper@59324
   829
      else if Context.proper_subthy_id p then LESS
wneuper@59324
   830
      else if Context.proper_subthy_id (swap p) then GREATER
wneuper@59324
   831
      else
wneuper@59324
   832
        (case apply2 ancestor_length p of
wneuper@59324
   833
          (SOME m, SOME n) =>
wneuper@59324
   834
            (case int_ord (m, n) of
wneuper@59324
   835
              EQUAL => string_ord (apply2 Context.theory_id_name p)
wneuper@59324
   836
            | ord => ord)
wneuper@59324
   837
        | _ => string_ord (apply2 Context.theory_id_name p))
wneuper@59324
   838
  in
wneuper@59324
   839
    fn p =>
wneuper@59324
   840
      (case crude_theory_ord (apply2 Thm.theory_id p) of
wneuper@59324
   841
        EQUAL =>
wneuper@59324
   842
        (* The hack below is necessary because of odd dependencies that are not reflected in the theory
wneuper@59324
   843
           comparison. *)
wneuper@59324
   844
        let val q = apply2 nickname_of_thm p in
wneuper@59324
   845
          (* Hack to put "xxx_def" before "xxxI" and "xxxE" *)
wneuper@59324
   846
          (case bool_ord (apply2 (String.isSuffix "_def") (swap q)) of
wneuper@59324
   847
            EQUAL => string_ord q
wneuper@59324
   848
          | ord => ord)
wneuper@59324
   849
        end
blanchet@56628
   850
      | ord => ord)
wneuper@59324
   851
  end;
blanchet@49339
   852
wneuper@59324
   853
val thm_less_eq = Context.subthy_id o apply2 Thm.theory_id
blanchet@52273
   854
fun thm_less p = thm_less_eq p andalso not (thm_less_eq (swap p))
blanchet@52273
   855
blanchet@49407
   856
val freezeT = Type.legacy_freeze_type
blanchet@49407
   857
blanchet@49407
   858
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49407
   859
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49407
   860
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49407
   861
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49407
   862
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49407
   863
  | freeze t = t
blanchet@49407
   864
wneuper@59180
   865
fun goal_of_thm thy = Thm.prop_of #> freeze #> Thm.global_cterm_of thy #> Goal.init
blanchet@49407
   866
blanchet@55593
   867
fun run_prover_for_mash ctxt params prover goal_name facts goal =
blanchet@49407
   868
  let
blanchet@49407
   869
    val problem =
blanchet@55593
   870
      {comment = "Goal: " ^ goal_name, state = Proof.init ctxt, goal = goal, subgoal = 1,
wneuper@59324
   871
       subgoal_count = 1, factss = [("", facts)], found_proof = I}
blanchet@49407
   872
  in
wneuper@59180
   873
    get_minimizing_prover ctxt MaSh (K ()) prover params problem
blanchet@49407
   874
  end
blanchet@49407
   875
walther@59606
   876
val bad_types = [\<^type_name>\<open>prop\<close>, \<^type_name>\<open>bool\<close>, \<^type_name>\<open>fun\<close>]
blanchet@49341
   877
walther@59606
   878
val crude_str_of_sort = space_implode "," o map Long_Name.base_name o subtract (op =) \<^sort>\<open>type\<close>
blanchet@54220
   879
wneuper@59324
   880
fun crude_str_of_typ (Type (s, [])) = Long_Name.base_name s
wneuper@59324
   881
  | crude_str_of_typ (Type (s, Ts)) = Long_Name.base_name s ^ implode (map crude_str_of_typ Ts)
blanchet@54220
   882
  | crude_str_of_typ (TFree (_, S)) = crude_str_of_sort S
blanchet@54220
   883
  | crude_str_of_typ (TVar (_, S)) = crude_str_of_sort S
blanchet@54220
   884
wneuper@59324
   885
fun maybe_singleton_str "" = []
wneuper@59324
   886
  | maybe_singleton_str s = [s]
blanchet@54265
   887
wneuper@59324
   888
val max_pat_breadth = 5 (* FUDGE *)
blanchet@51600
   889
blanchet@58748
   890
fun term_features_of ctxt thy_name term_max_depth type_max_depth ts =
blanchet@49266
   891
  let
blanchet@51407
   892
    val thy = Proof_Context.theory_of ctxt
blanchet@54219
   893
blanchet@51408
   894
    val fixes = map snd (Variable.dest_fixes ctxt)
blanchet@51407
   895
    val classes = Sign.classes_of thy
blanchet@54219
   896
walther@59606
   897
    fun add_classes \<^sort>\<open>type\<close> = I
blanchet@51407
   898
      | add_classes S =
blanchet@51407
   899
        fold (`(Sorts.super_classes classes)
blanchet@58348
   900
          #> swap #> op ::
walther@59606
   901
          #> subtract (op =) \<^sort>\<open>type\<close>
blanchet@58348
   902
          #> map class_feature_of
blanchet@58746
   903
          #> union (op =)) S
blanchet@54219
   904
blanchet@54219
   905
    fun pattify_type 0 _ = []
walther@59606
   906
      | pattify_type _ (Type (s, [])) = if member (op =) bad_types s then [] else [s]
blanchet@54219
   907
      | pattify_type depth (Type (s, U :: Ts)) =
blanchet@54219
   908
        let
blanchet@54219
   909
          val T = Type (s, Ts)
blanchet@58397
   910
          val ps = take max_pat_breadth (pattify_type depth T)
blanchet@58397
   911
          val qs = take max_pat_breadth ("" :: pattify_type (depth - 1) U)
blanchet@58348
   912
        in
blanchet@58348
   913
          map_product (fn p => fn "" => p | q => p ^ "(" ^ q ^ ")") ps qs
blanchet@58348
   914
        end
wneuper@59324
   915
      | pattify_type _ (TFree (_, S)) = maybe_singleton_str (crude_str_of_sort S)
wneuper@59324
   916
      | pattify_type _ (TVar (_, S)) = maybe_singleton_str (crude_str_of_sort S)
blanchet@58348
   917
blanchet@54219
   918
    fun add_type_pat depth T =
blanchet@58746
   919
      union (op =) (map type_feature_of (pattify_type depth T))
blanchet@58348
   920
blanchet@54219
   921
    fun add_type_pats 0 _ = I
blanchet@58746
   922
      | add_type_pats depth t = add_type_pat depth t #> add_type_pats (depth - 1) t
blanchet@58348
   923
blanchet@54220
   924
    fun add_type T =
blanchet@54220
   925
      add_type_pats type_max_depth T
blanchet@54293
   926
      #> fold_atyps_sorts (add_classes o snd) T
blanchet@58348
   927
blanchet@54221
   928
    fun add_subtypes (T as Type (_, Ts)) = add_type T #> fold add_subtypes Ts
blanchet@54221
   929
      | add_subtypes T = add_type T
blanchet@54219
   930
blanchet@58397
   931
    fun pattify_term _ 0 _ = []
walther@59606
   932
      | pattify_term _ _ (Const (s, _)) =
wneuper@59324
   933
        if is_widely_irrelevant_const s then [] else [s]
blanchet@55541
   934
      | pattify_term _ _ (Free (s, T)) =
wneuper@59324
   935
        maybe_singleton_str (crude_str_of_typ T)
wneuper@59324
   936
        |> (if member (op =) fixes s then cons (free_feature_of (Long_Name.append thy_name s))
wneuper@59324
   937
            else I)
wneuper@59324
   938
      | pattify_term _ _ (Var (_, T)) =
wneuper@59324
   939
        maybe_singleton_str (crude_str_of_typ T)
blanchet@55541
   940
      | pattify_term Ts _ (Bound j) =
wneuper@59324
   941
        maybe_singleton_str (crude_str_of_typ (nth Ts j))
blanchet@55541
   942
      | pattify_term Ts depth (t $ u) =
blanchet@51354
   943
        let
blanchet@58397
   944
          val ps = take max_pat_breadth (pattify_term Ts depth t)
blanchet@58746
   945
          val qs = take max_pat_breadth ("" :: pattify_term Ts (depth - 1) u)
blanchet@54267
   946
        in
blanchet@58746
   947
          map_product (fn p => fn "" => p | q => p ^ "(" ^ q ^ ")") ps qs
blanchet@54267
   948
        end
blanchet@55541
   949
      | pattify_term _ _ _ = []
blanchet@58348
   950
blanchet@58746
   951
    fun add_term_pat Ts = union (op =) oo pattify_term Ts
blanchet@58348
   952
blanchet@54267
   953
    fun add_term_pats _ 0 _ = I
blanchet@58397
   954
      | add_term_pats Ts depth t = add_term_pat Ts depth t #> add_term_pats Ts (depth - 1) t
blanchet@58348
   955
blanchet@54267
   956
    fun add_term Ts = add_term_pats Ts term_max_depth
blanchet@58348
   957
blanchet@54222
   958
    fun add_subterms Ts t =
blanchet@56628
   959
      (case strip_comb t of
blanchet@55541
   960
        (Const (s, T), args) =>
blanchet@55541
   961
        (not (is_widely_irrelevant_const s) ? add_term Ts t)
blanchet@58746
   962
        #> add_subtypes T #> fold (add_subterms Ts) args
blanchet@51872
   963
      | (head, args) =>
blanchet@49266
   964
        (case head of
blanchet@54267
   965
           Free (_, T) => add_term Ts t #> add_subtypes T
blanchet@54221
   966
         | Var (_, T) => add_subtypes T
blanchet@54222
   967
         | Abs (_, T, body) => add_subtypes T #> add_subterms (T :: Ts) body
blanchet@49266
   968
         | _ => I)
blanchet@56628
   969
        #> fold (add_subterms Ts) args)
blanchet@58348
   970
  in
blanchet@58348
   971
    fold (add_subterms []) ts []
blanchet@58348
   972
  end
blanchet@49266
   973
blanchet@54222
   974
val term_max_depth = 2
blanchet@54292
   975
val type_max_depth = 1
blanchet@49266
   976
blanchet@49266
   977
(* TODO: Generate type classes for types? *)
wneuper@59324
   978
fun features_of ctxt thy_name (scope, _) ts =
wneuper@59324
   979
  thy_feature_of thy_name ::
wneuper@59324
   980
  term_features_of ctxt thy_name term_max_depth type_max_depth ts
wneuper@59324
   981
  |> scope <> Global ? cons local_feature
blanchet@49266
   982
blanchet@58348
   983
(* Too many dependencies is a sign that a decision procedure is at work. There is not much to learn
blanchet@58348
   984
   from such proofs. *)
wneuper@59324
   985
val max_dependencies = 20 (* FUDGE *)
blanchet@51499
   986
wneuper@59324
   987
val prover_default_max_facts = 25 (* FUDGE *)
blanchet@49266
   988
blanchet@49453
   989
(* "type_definition_xxx" facts are characterized by their use of "CollectI". *)
blanchet@51770
   990
val typedef_dep = nickname_of_thm @{thm CollectI}
blanchet@58348
   991
(* Mysterious parts of the class machinery create lots of proofs that refer exclusively to
blanchet@58348
   992
   "someI_ex" (and to some internal constructions). *)
blanchet@51770
   993
val class_some_dep = nickname_of_thm @{thm someI_ex}
blanchet@49453
   994
blanchet@51843
   995
val fundef_ths =
blanchet@58348
   996
  @{thms fundef_ex1_existence fundef_ex1_uniqueness fundef_ex1_iff fundef_default_value}
blanchet@51843
   997
  |> map nickname_of_thm
blanchet@51843
   998
blanchet@49453
   999
(* "Rep_xxx_inject", "Abs_xxx_inverse", etc., are derived using these facts. *)
blanchet@49453
  1000
val typedef_ths =
blanchet@58348
  1001
  @{thms type_definition.Abs_inverse type_definition.Rep_inverse type_definition.Rep
blanchet@58348
  1002
      type_definition.Rep_inject type_definition.Abs_inject type_definition.Rep_cases
blanchet@58348
  1003
      type_definition.Abs_cases type_definition.Rep_induct type_definition.Abs_induct
blanchet@58348
  1004
      type_definition.Rep_range type_definition.Abs_image}
blanchet@51639
  1005
  |> map nickname_of_thm
blanchet@49453
  1006
blanchet@49456
  1007
fun is_size_def [dep] th =
blanchet@56984
  1008
    (case first_field ".rec" dep of
blanchet@58348
  1009
      SOME (pref, _) =>
blanchet@58348
  1010
      (case first_field ".size" (nickname_of_thm th) of
blanchet@58348
  1011
        SOME (pref', _) => pref = pref'
blanchet@58348
  1012
      | NONE => false)
blanchet@58348
  1013
    | NONE => false)
blanchet@49456
  1014
  | is_size_def _ _ = false
blanchet@49456
  1015
blanchet@52314
  1016
fun trim_dependencies deps =
blanchet@51770
  1017
  if length deps > max_dependencies then NONE else SOME deps
blanchet@49266
  1018
blanchet@51770
  1019
fun isar_dependencies_of name_tabs th =
blanchet@58648
  1020
  thms_in_proof max_dependencies (SOME name_tabs) th
blanchet@58648
  1021
  |> Option.map (fn deps =>
blanchet@58348
  1022
    if deps = [typedef_dep] orelse deps = [class_some_dep] orelse
wneuper@59180
  1023
        exists (member (op =) fundef_ths) deps orelse exists (member (op =) typedef_ths) deps orelse
wneuper@59180
  1024
        is_size_def deps th then
blanchet@51770
  1025
      []
blanchet@51770
  1026
    else
blanchet@58648
  1027
      deps)
blanchet@49419
  1028
blanchet@58348
  1029
fun prover_dependencies_of ctxt (params as {verbose, max_facts, ...}) prover auto_level facts
blanchet@58348
  1030
    name_tabs th =
blanchet@56628
  1031
  (case isar_dependencies_of name_tabs th of
blanchet@58648
  1032
    SOME [] => (false, [])
blanchet@58648
  1033
  | isar_deps0 =>
blanchet@49407
  1034
    let
blanchet@58648
  1035
      val isar_deps = these isar_deps0
blanchet@49407
  1036
      val thy = Proof_Context.theory_of ctxt
blanchet@49407
  1037
      val goal = goal_of_thm thy th
blanchet@55593
  1038
      val name = nickname_of_thm th
blanchet@53333
  1039
      val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal 1 ctxt
blanchet@52273
  1040
      val facts = facts |> filter (fn (_, th') => thm_less (th', th))
blanchet@58348
  1041
blanchet@52186
  1042
      fun nickify ((_, stature), th) = ((nickname_of_thm th, stature), th)
blanchet@58348
  1043
wneuper@59324
  1044
      fun is_dep dep (_, th) = (nickname_of_thm th = dep)
blanchet@58348
  1045
blanchet@49407
  1046
      fun add_isar_dep facts dep accum =
blanchet@49407
  1047
        if exists (is_dep dep) accum then
blanchet@49407
  1048
          accum
blanchet@56628
  1049
        else
blanchet@56628
  1050
          (case find_first (is_dep dep) facts of
blanchet@56628
  1051
            SOME ((_, status), th) => accum @ [(("", status), th)]
blanchet@58347
  1052
          | NONE => accum (* should not happen *))
blanchet@58348
  1053
blanchet@55575
  1054
      val mepo_facts =
blanchet@51499
  1055
        facts
blanchet@55547
  1056
        |> mepo_suggested_facts ctxt params (max_facts |> the_default prover_default_max_facts) NONE
blanchet@55547
  1057
             hyp_ts concl_t
blanchet@55575
  1058
      val facts =
blanchet@55575
  1059
        mepo_facts
blanchet@51769
  1060
        |> fold (add_isar_dep facts) isar_deps
blanchet@51639
  1061
        |> map nickify
blanchet@55575
  1062
      val num_isar_deps = length isar_deps
blanchet@49407
  1063
    in
blanchet@49419
  1064
      if verbose andalso auto_level = 0 then
wneuper@59180
  1065
        writeln ("MaSh: " ^ quote prover ^ " on " ^ quote name ^ " with " ^
blanchet@58359
  1066
          string_of_int num_isar_deps ^ " + " ^ string_of_int (length facts - num_isar_deps) ^
wneuper@59324
  1067
          " facts")
blanchet@49407
  1068
      else
blanchet@49407
  1069
        ();
blanchet@56628
  1070
      (case run_prover_for_mash ctxt params prover name facts goal of
blanchet@49407
  1071
        {outcome = NONE, used_facts, ...} =>
blanchet@49419
  1072
        (if verbose andalso auto_level = 0 then
blanchet@49407
  1073
           let val num_facts = length used_facts in
wneuper@59180
  1074
             writeln ("Found proof with " ^ string_of_int num_facts ^ " fact" ^
wneuper@59324
  1075
               plural_s num_facts)
blanchet@49407
  1076
           end
blanchet@49407
  1077
         else
blanchet@49407
  1078
           ();
blanchet@51769
  1079
         (true, map fst used_facts))
blanchet@56628
  1080
      | _ => (false, isar_deps))
blanchet@56628
  1081
    end)
blanchet@49266
  1082
blanchet@49266
  1083
blanchet@49266
  1084
(*** High-level communication with MaSh ***)
blanchet@49266
  1085
blanchet@52319
  1086
(* In the following functions, chunks are risers w.r.t. "thm_less_eq". *)
blanchet@52319
  1087
blanchet@52318
  1088
fun chunks_and_parents_for chunks th =
blanchet@52318
  1089
  let
blanchet@52318
  1090
    fun insert_parent new parents =
blanchet@52318
  1091
      let val parents = parents |> filter_out (fn p => thm_less_eq (p, new)) in
blanchet@58438
  1092
        parents |> forall (fn p => not (thm_less_eq (new, p))) parents ? cons new
blanchet@52318
  1093
      end
blanchet@58348
  1094
blanchet@52318
  1095
    fun rechunk seen (rest as th' :: ths) =
blanchet@52318
  1096
      if thm_less_eq (th', th) then (rev seen, rest)
blanchet@52318
  1097
      else rechunk (th' :: seen) ths
blanchet@58348
  1098
blanchet@52318
  1099
    fun do_chunk [] accum = accum
blanchet@52318
  1100
      | do_chunk (chunk as hd_chunk :: _) (chunks, parents) =
blanchet@52318
  1101
        if thm_less_eq (hd_chunk, th) then
blanchet@52318
  1102
          (chunk :: chunks, insert_parent hd_chunk parents)
blanchet@52318
  1103
        else if thm_less_eq (List.last chunk, th) then
blanchet@52318
  1104
          let val (front, back as hd_back :: _) = rechunk [] chunk in
blanchet@52318
  1105
            (front :: back :: chunks, insert_parent hd_back parents)
blanchet@52318
  1106
          end
blanchet@52318
  1107
        else
blanchet@52318
  1108
          (chunk :: chunks, parents)
blanchet@52318
  1109
  in
blanchet@52318
  1110
    fold_rev do_chunk chunks ([], [])
blanchet@52318
  1111
    |>> cons []
blanchet@52319
  1112
    ||> map nickname_of_thm
blanchet@52318
  1113
  end
blanchet@52318
  1114
blanchet@52319
  1115
fun attach_parents_to_facts _ [] = []
blanchet@52319
  1116
  | attach_parents_to_facts old_facts (facts as (_, th) :: _) =
blanchet@52319
  1117
    let
blanchet@52319
  1118
      fun do_facts _ [] = []
blanchet@52319
  1119
        | do_facts (_, parents) [fact] = [(parents, fact)]
blanchet@52319
  1120
        | do_facts (chunks, parents)
blanchet@52319
  1121
                   ((fact as (_, th)) :: (facts as (_, th') :: _)) =
blanchet@52319
  1122
          let
blanchet@52319
  1123
            val chunks = app_hd (cons th) chunks
blanchet@52319
  1124
            val chunks_and_parents' =
wneuper@59324
  1125
              if thm_less_eq (th, th') andalso
wneuper@59324
  1126
                Thm.theory_name th = Thm.theory_name th'
wneuper@59324
  1127
              then (chunks, [nickname_of_thm th])
wneuper@59324
  1128
              else chunks_and_parents_for chunks th'
blanchet@58348
  1129
          in
blanchet@58348
  1130
            (parents, fact) :: do_facts chunks_and_parents' facts
blanchet@58348
  1131
          end
blanchet@52319
  1132
    in
blanchet@52319
  1133
      old_facts @ facts
blanchet@52319
  1134
      |> do_facts (chunks_and_parents_for [[]] th)
blanchet@52319
  1135
      |> drop (length old_facts)
blanchet@52319
  1136
    end
blanchet@52314
  1137
blanchet@54232
  1138
fun is_fact_in_graph access_G = can (Graph.get_node access_G) o nickname_of_thm
blanchet@54232
  1139
blanchet@54334
  1140
val chained_feature_factor = 0.5 (* FUDGE *)
blanchet@58747
  1141
val extra_feature_factor = 0.1 (* FUDGE *)
blanchet@58747
  1142
val num_extra_feature_facts = 10 (* FUDGE *)
blanchet@54232
  1143
blanchet@59002
  1144
val max_proximity_facts = 100 (* FUDGE *)
blanchet@54232
  1145
blanchet@55512
  1146
fun find_mash_suggestions ctxt max_facts suggs facts chained raw_unknown =
blanchet@55512
  1147
  let
blanchet@55512
  1148
    val inter_fact = inter (eq_snd Thm.eq_thm_prop)
blanchet@55512
  1149
    val raw_mash = find_suggested_facts ctxt facts suggs
blanchet@55512
  1150
    val proximate = take max_proximity_facts facts
blanchet@55512
  1151
    val unknown_chained = inter_fact raw_unknown chained
blanchet@55512
  1152
    val unknown_proximate = inter_fact raw_unknown proximate
blanchet@55512
  1153
    val mess =
blanchet@55512
  1154
      [(0.9 (* FUDGE *), (map (rpair 1.0) unknown_chained, [])),
blanchet@55512
  1155
       (0.4 (* FUDGE *), (weight_facts_smoothly unknown_proximate, [])),
blanchet@55512
  1156
       (0.1 (* FUDGE *), (weight_facts_steeply raw_mash, raw_unknown))]
blanchet@58349
  1157
    val unknown = raw_unknown
blanchet@58349
  1158
      |> fold (subtract (eq_snd Thm.eq_thm_prop)) [unknown_chained, unknown_proximate]
blanchet@58348
  1159
  in
wneuper@59324
  1160
    (mesh_facts (fact_distinct (op aconv)) (eq_snd (gen_eq_thm ctxt)) max_facts mess, unknown)
blanchet@58348
  1161
  end
blanchet@54232
  1162
wneuper@59324
  1163
fun mash_suggested_facts ctxt thy_name ({debug, ...} : params) max_suggs hyp_ts concl_t facts =
blanchet@54232
  1164
  let
blanchet@58874
  1165
    val algorithm = the_mash_algorithm ()
blanchet@58394
  1166
blanchet@58802
  1167
    val facts = facts
wneuper@59324
  1168
      |> rev_sort_list_prefix (crude_thm_ord ctxt o apply2 snd)
blanchet@58802
  1169
        (Int.max (num_extra_feature_facts, max_proximity_facts))
blanchet@58802
  1170
blanchet@58802
  1171
    val chained = filter (fn ((_, (scope, _)), _) => scope = Chained) facts
blanchet@55537
  1172
wneuper@59324
  1173
    fun fact_has_right_theory (_, th) = thy_name = Thm.theory_name th
blanchet@58337
  1174
blanchet@54278
  1175
    fun chained_or_extra_features_of factor (((_, stature), th), weight) =
wneuper@59180
  1176
      [Thm.prop_of th]
wneuper@59324
  1177
      |> features_of ctxt (Thm.theory_name th) stature
blanchet@58746
  1178
      |> map (rpair (weight * factor))
wneuper@59324
  1179
  in
wneuper@59324
  1180
    (case get_state ctxt of
wneuper@59324
  1181
      NONE => ([], [])
wneuper@59324
  1182
    | SOME {access_G, xtabs = ((num_facts, fact_tab), (num_feats, feat_tab)), ffds, freqs, ...} =>
wneuper@59324
  1183
      let
wneuper@59324
  1184
        val goal_feats0 =
wneuper@59324
  1185
          features_of ctxt thy_name (Local, General) (concl_t :: hyp_ts)
wneuper@59324
  1186
        val chained_feats = chained
wneuper@59324
  1187
          |> map (rpair 1.0)
wneuper@59324
  1188
          |> map (chained_or_extra_features_of chained_feature_factor)
wneuper@59324
  1189
          |> rpair [] |-> fold (union (eq_fst (op =)))
wneuper@59324
  1190
        val extra_feats = facts
wneuper@59324
  1191
          |> take (Int.max (0, num_extra_feature_facts - length chained))
wneuper@59324
  1192
          |> filter fact_has_right_theory
wneuper@59324
  1193
          |> weight_facts_steeply
wneuper@59324
  1194
          |> map (chained_or_extra_features_of extra_feature_factor)
wneuper@59324
  1195
          |> rpair [] |-> fold (union (eq_fst (op =)))
blanchet@55537
  1196
wneuper@59324
  1197
        val goal_feats =
wneuper@59324
  1198
          fold (union (eq_fst (op =))) [chained_feats, extra_feats] (map (rpair 1.0) goal_feats0)
wneuper@59324
  1199
          |> debug ? sort (Real.compare o swap o apply2 snd)
blanchet@58359
  1200
walther@59606
  1201
        val fact_idxs = map_filter (Symtab.lookup fact_tab o nickname_of_thm o snd) facts
blanchet@58802
  1202
wneuper@59324
  1203
        val suggs =
wneuper@59324
  1204
          if algorithm = MaSh_NB_Ext orelse algorithm = MaSh_kNN_Ext then
wneuper@59324
  1205
            let
wneuper@59324
  1206
              val learns =
wneuper@59324
  1207
                Graph.schedule (fn _ => fn (fact, (_, feats, deps)) => (fact, feats, deps))
wneuper@59324
  1208
                  access_G
wneuper@59324
  1209
            in
wneuper@59324
  1210
              MaSh.query_external ctxt algorithm max_suggs learns goal_feats
wneuper@59324
  1211
            end
wneuper@59324
  1212
          else
wneuper@59324
  1213
            let
wneuper@59324
  1214
              val int_goal_feats =
wneuper@59324
  1215
                map_filter (fn (s, w) => Option.map (rpair w) (Symtab.lookup feat_tab s)) goal_feats
wneuper@59324
  1216
            in
walther@59606
  1217
              MaSh.query_internal ctxt algorithm num_facts num_feats ffds freqs fact_idxs max_suggs
walther@59606
  1218
                goal_feats int_goal_feats
wneuper@59324
  1219
            end
blanchet@58802
  1220
wneuper@59324
  1221
        val unknown = filter_out (is_fact_in_graph access_G o snd) facts
wneuper@59324
  1222
      in
wneuper@59324
  1223
        find_mash_suggestions ctxt max_suggs suggs facts chained unknown
wneuper@59324
  1224
        |> apply2 (map fact_of_raw_fact)
wneuper@59324
  1225
      end)
blanchet@54232
  1226
  end
blanchet@54232
  1227
wneuper@59324
  1228
fun mash_unlearn ctxt = (clear_state ctxt; writeln "Reset MaSh")
blanchet@58773
  1229
blanchet@59003
  1230
fun learn_wrt_access_graph ctxt (name, parents, feats, deps)
blanchet@59003
  1231
    (accum as (access_G, (fact_xtab, feat_xtab))) =
blanchet@54232
  1232
  let
blanchet@58713
  1233
    fun maybe_learn_from from (accum as (parents, access_G)) =
blanchet@58355
  1234
      try_graph ctxt "updating graph" accum (fn () =>
blanchet@58713
  1235
        (from :: parents, Graph.add_edge_acyclic (from, name) access_G))
blanchet@58713
  1236
blanchet@58713
  1237
    val access_G = access_G |> Graph.default_node (name, (Isar_Proof, feats, deps))
blanchet@58713
  1238
    val (parents, access_G) = ([], access_G) |> fold maybe_learn_from parents
blanchet@58713
  1239
    val (deps, _) = ([], access_G) |> fold maybe_learn_from deps
blanchet@58713
  1240
blanchet@59003
  1241
    val fact_xtab = add_to_xtab name fact_xtab
blanchet@58713
  1242
    val feat_xtab = fold maybe_add_to_xtab feats feat_xtab
blanchet@58348
  1243
  in
blanchet@59003
  1244
    (SOME (name, parents, feats, deps), (access_G, (fact_xtab, feat_xtab)))
blanchet@58348
  1245
  end
blanchet@59003
  1246
  handle Symtab.DUP _ => (NONE, accum) (* facts sometimes have the same name, confusingly *)
blanchet@54232
  1247
blanchet@58723
  1248
fun relearn_wrt_access_graph ctxt (name, deps) access_G =
blanchet@54232
  1249
  let
blanchet@58713
  1250
    fun maybe_relearn_from from (accum as (parents, access_G)) =
blanchet@54232
  1251
      try_graph ctxt "updating graph" accum (fn () =>
blanchet@58713
  1252
        (from :: parents, Graph.add_edge_acyclic (from, name) access_G))
blanchet@58713
  1253
    val access_G =
blanchet@58713
  1254
      access_G |> Graph.map_node name (fn (_, feats, _) => (Automatic_Proof, feats, deps))
blanchet@58713
  1255
    val (deps, _) = ([], access_G) |> fold maybe_relearn_from deps
blanchet@58348
  1256
  in
blanchet@58723
  1257
    ((name, deps), access_G)
blanchet@58348
  1258
  end
blanchet@54232
  1259
blanchet@54232
  1260
fun flop_wrt_access_graph name =
blanchet@58347
  1261
  Graph.map_node name (fn (_, feats, deps) => (Isar_Proof_wegen_Prover_Flop, feats, deps))
blanchet@54232
  1262
blanchet@58615
  1263
val learn_timeout_slack = 20.0
blanchet@54232
  1264
blanchet@54232
  1265
fun launch_thread timeout task =
blanchet@54232
  1266
  let
blanchet@54232
  1267
    val hard_timeout = time_mult learn_timeout_slack timeout
blanchet@54232
  1268
    val birth_time = Time.now ()
wneuper@59324
  1269
    val death_time = birth_time + hard_timeout
blanchet@54232
  1270
    val desc = ("Machine learner for Sledgehammer", "")
blanchet@58348
  1271
  in
wneuper@59180
  1272
    Async_Manager_Legacy.thread MaShN birth_time death_time desc task
blanchet@58348
  1273
  end
blanchet@54232
  1274
wneuper@59324
  1275
fun anonymous_proof_name () =
wneuper@59324
  1276
  Date.fmt (anonymous_proof_prefix ^ "%Y%m%d.%H%M%S.") (Date.fromTimeLocal (Time.now ())) ^
wneuper@59324
  1277
  serial_string ()
blanchet@58355
  1278
walther@59606
  1279
fun mash_learn_proof ctxt ({timeout, ...} : params) t used_ths =
blanchet@58729
  1280
  if not (null used_ths) andalso is_mash_enabled () then
blanchet@56158
  1281
    launch_thread timeout (fn () =>
blanchet@58348
  1282
      let
blanchet@58348
  1283
        val thy = Proof_Context.theory_of ctxt
wneuper@59324
  1284
        val feats = features_of ctxt (Context.theory_name thy) (Local, General) [t]
blanchet@58348
  1285
      in
blanchet@58773
  1286
        map_state ctxt
blanchet@58773
  1287
          (fn {access_G, xtabs as ((num_facts0, _), _), ffds, freqs, dirty_facts} =>
blanchet@58713
  1288
             let
blanchet@58713
  1289
               val deps = used_ths
blanchet@58713
  1290
                 |> filter (is_fact_in_graph access_G)
blanchet@58713
  1291
                 |> map nickname_of_thm
blanchet@58773
  1292
wneuper@59324
  1293
               val name = anonymous_proof_name ()
blanchet@58773
  1294
               val (access_G', xtabs', rev_learns) =
walther@59606
  1295
                 add_node Automatic_Proof name [] (* ignore parents *) feats deps
walther@59606
  1296
                   (access_G, xtabs, [])
blanchet@58773
  1297
blanchet@58773
  1298
               val (ffds', freqs') =
blanchet@58773
  1299
                 recompute_ffds_freqs_from_learns (rev rev_learns) xtabs' num_facts0 ffds freqs
blanchet@58713
  1300
             in
blanchet@58773
  1301
               {access_G = access_G', xtabs = xtabs', ffds = ffds', freqs = freqs',
blanchet@58773
  1302
                dirty_facts = Option.map (cons name) dirty_facts}
blanchet@58713
  1303
             end);
blanchet@58348
  1304
        (true, "")
blanchet@58348
  1305
      end)
blanchet@54956
  1306
  else
blanchet@54956
  1307
    ()
blanchet@54232
  1308
wneuper@59324
  1309
fun sendback sub = Active.sendback_markup_command (sledgehammerN ^ " " ^ sub)
blanchet@49407
  1310
blanchet@49407
  1311
val commit_timeout = seconds 30.0
blanchet@49347
  1312
blanchet@51500
  1313
(* The timeout is understood in a very relaxed fashion. *)
blanchet@58773
  1314
fun mash_learn_facts ctxt (params as {debug, verbose, ...}) prover auto_level run_prover
blanchet@58773
  1315
    learn_timeout facts =
blanchet@49319
  1316
  let
blanchet@49333
  1317
    val timer = Timer.startRealTimer ()
wneuper@59324
  1318
    fun next_commit_time () = Timer.checkRealTimer timer + commit_timeout
wneuper@59324
  1319
  in
wneuper@59324
  1320
    (case get_state ctxt of
wneuper@59324
  1321
      NONE => "MaSh is busy\nPlease try again later"
wneuper@59324
  1322
    | SOME {access_G, ...} =>
wneuper@59324
  1323
      let
wneuper@59324
  1324
        val is_in_access_G = is_fact_in_graph access_G o snd
wneuper@59324
  1325
        val no_new_facts = forall is_in_access_G facts
wneuper@59324
  1326
      in
wneuper@59324
  1327
        if no_new_facts andalso not run_prover then
wneuper@59324
  1328
          if auto_level < 2 then
wneuper@59324
  1329
            "No new " ^ (if run_prover then "automatic" else "Isar") ^ " proofs to learn" ^
wneuper@59324
  1330
            (if auto_level = 0 andalso not run_prover then
wneuper@59324
  1331
               "\n\nHint: Try " ^ sendback learn_proverN ^ " to learn from an automatic prover"
wneuper@59324
  1332
             else
wneuper@59324
  1333
               "")
wneuper@59324
  1334
          else
wneuper@59324
  1335
            ""
wneuper@59324
  1336
        else
wneuper@59324
  1337
          let
wneuper@59324
  1338
            val name_tabs = build_name_tables nickname_of_thm facts
blanchet@58347
  1339
wneuper@59324
  1340
            fun deps_of status th =
wneuper@59324
  1341
              if status = Non_Rec_Def orelse status = Rec_Def then
wneuper@59324
  1342
                SOME []
wneuper@59324
  1343
              else if run_prover then
wneuper@59324
  1344
                prover_dependencies_of ctxt params prover auto_level facts name_tabs th
wneuper@59324
  1345
                |> (fn (false, _) => NONE | (true, deps) => trim_dependencies deps)
wneuper@59324
  1346
              else
wneuper@59324
  1347
                isar_dependencies_of name_tabs th
blanchet@58347
  1348
wneuper@59324
  1349
            fun do_commit [] [] [] state = state
wneuper@59324
  1350
              | do_commit learns relearns flops
wneuper@59324
  1351
                  {access_G, xtabs as ((num_facts0, _), _), ffds, freqs, dirty_facts} =
wneuper@59324
  1352
                let
wneuper@59324
  1353
                  val was_empty = Graph.is_empty access_G
blanchet@58347
  1354
wneuper@59324
  1355
                  val (learns, (access_G', xtabs')) =
wneuper@59324
  1356
                    fold_map (learn_wrt_access_graph ctxt) learns (access_G, xtabs)
wneuper@59324
  1357
                    |>> map_filter I
wneuper@59324
  1358
                  val (relearns, access_G'') =
wneuper@59324
  1359
                    fold_map (relearn_wrt_access_graph ctxt) relearns access_G'
blanchet@58720
  1360
wneuper@59324
  1361
                  val access_G''' = access_G'' |> fold flop_wrt_access_graph flops
wneuper@59324
  1362
                  val dirty_facts' =
wneuper@59324
  1363
                    (case (was_empty, dirty_facts) of
wneuper@59324
  1364
                      (false, SOME names) => SOME (map #1 learns @ map #1 relearns @ names)
wneuper@59324
  1365
                    | _ => NONE)
blanchet@58713
  1366
wneuper@59324
  1367
                  val (ffds', freqs') =
wneuper@59324
  1368
                    if null relearns then
wneuper@59324
  1369
                      recompute_ffds_freqs_from_learns
wneuper@59324
  1370
                        (map (fn (name, _, feats, deps) => (name, feats, deps)) learns) xtabs'
wneuper@59324
  1371
                        num_facts0 ffds freqs
wneuper@59324
  1372
                    else
wneuper@59324
  1373
                      recompute_ffds_freqs_from_access_G access_G''' xtabs'
wneuper@59324
  1374
                in
wneuper@59324
  1375
                  {access_G = access_G''', xtabs = xtabs', ffds = ffds', freqs = freqs',
wneuper@59324
  1376
                   dirty_facts = dirty_facts'}
wneuper@59324
  1377
                end
blanchet@58720
  1378
wneuper@59324
  1379
            fun commit last learns relearns flops =
wneuper@59324
  1380
              (if debug andalso auto_level = 0 then writeln "Committing..." else ();
wneuper@59324
  1381
               map_state ctxt (do_commit (rev learns) relearns flops);
wneuper@59324
  1382
               if not last andalso auto_level = 0 then
wneuper@59324
  1383
                 let val num_proofs = length learns + length relearns in
wneuper@59324
  1384
                   writeln ("Learned " ^ string_of_int num_proofs ^ " " ^
wneuper@59324
  1385
                     (if run_prover then "automatic" else "Isar") ^ " proof" ^
wneuper@59324
  1386
                     plural_s num_proofs ^ " in the last " ^ string_of_time commit_timeout)
wneuper@59324
  1387
                 end
wneuper@59324
  1388
               else
wneuper@59324
  1389
                 ())
blanchet@58347
  1390
wneuper@59324
  1391
            fun learn_new_fact _ (accum as (_, (_, _, true))) = accum
wneuper@59324
  1392
              | learn_new_fact (parents, ((_, stature as (_, status)), th))
wneuper@59324
  1393
                  (learns, (num_nontrivial, next_commit, _)) =
wneuper@59324
  1394
                let
wneuper@59324
  1395
                  val name = nickname_of_thm th
wneuper@59324
  1396
                  val feats = features_of ctxt (Thm.theory_name th) stature [Thm.prop_of th]
wneuper@59324
  1397
                  val deps = these (deps_of status th)
wneuper@59324
  1398
                  val num_nontrivial = num_nontrivial |> not (null deps) ? Integer.add 1
wneuper@59324
  1399
                  val learns = (name, parents, feats, deps) :: learns
wneuper@59324
  1400
                  val (learns, next_commit) =
wneuper@59324
  1401
                    if Timer.checkRealTimer timer > next_commit then
wneuper@59324
  1402
                      (commit false learns [] []; ([], next_commit_time ()))
wneuper@59324
  1403
                    else
wneuper@59324
  1404
                      (learns, next_commit)
wneuper@59324
  1405
                  val timed_out = Timer.checkRealTimer timer > learn_timeout
wneuper@59324
  1406
                in
wneuper@59324
  1407
                  (learns, (num_nontrivial, next_commit, timed_out))
wneuper@59324
  1408
                end
blanchet@58347
  1409
wneuper@59324
  1410
            val (num_new_facts, num_nontrivial) =
wneuper@59324
  1411
              if no_new_facts then
wneuper@59324
  1412
                (0, 0)
wneuper@59324
  1413
              else
wneuper@59324
  1414
                let
wneuper@59324
  1415
                  val new_facts = facts
wneuper@59324
  1416
                    |> sort (crude_thm_ord ctxt o apply2 snd)
walther@59606
  1417
                    |> map (pair []) (* ignore parents *)
wneuper@59324
  1418
                    |> filter_out (is_in_access_G o snd)
wneuper@59324
  1419
                  val (learns, (num_nontrivial, _, _)) =
wneuper@59324
  1420
                    ([], (0, next_commit_time (), false))
wneuper@59324
  1421
                    |> fold learn_new_fact new_facts
wneuper@59324
  1422
                in
wneuper@59324
  1423
                  commit true learns [] []; (length new_facts, num_nontrivial)
wneuper@59324
  1424
                end
blanchet@58347
  1425
wneuper@59324
  1426
            fun relearn_old_fact _ (accum as (_, (_, _, true))) = accum
wneuper@59324
  1427
              | relearn_old_fact ((_, (_, status)), th)
wneuper@59324
  1428
                  ((relearns, flops), (num_nontrivial, next_commit, _)) =
wneuper@59324
  1429
                let
wneuper@59324
  1430
                  val name = nickname_of_thm th
wneuper@59324
  1431
                  val (num_nontrivial, relearns, flops) =
wneuper@59324
  1432
                    (case deps_of status th of
wneuper@59324
  1433
                      SOME deps => (num_nontrivial + 1, (name, deps) :: relearns, flops)
wneuper@59324
  1434
                    | NONE => (num_nontrivial, relearns, name :: flops))
wneuper@59324
  1435
                  val (relearns, flops, next_commit) =
wneuper@59324
  1436
                    if Timer.checkRealTimer timer > next_commit then
wneuper@59324
  1437
                      (commit false [] relearns flops; ([], [], next_commit_time ()))
wneuper@59324
  1438
                    else
wneuper@59324
  1439
                      (relearns, flops, next_commit)
wneuper@59324
  1440
                  val timed_out = Timer.checkRealTimer timer > learn_timeout
wneuper@59324
  1441
                in
wneuper@59324
  1442
                  ((relearns, flops), (num_nontrivial, next_commit, timed_out))
wneuper@59324
  1443
                end
blanchet@58347
  1444
wneuper@59324
  1445
            val num_nontrivial =
wneuper@59324
  1446
              if not run_prover then
wneuper@59324
  1447
                num_nontrivial
wneuper@59324
  1448
              else
wneuper@59324
  1449
                let
wneuper@59324
  1450
                  val max_isar = 1000 * max_dependencies
blanchet@58347
  1451
wneuper@59324
  1452
                  fun priority_of th =
wneuper@59324
  1453
                    Random.random_range 0 max_isar +
wneuper@59324
  1454
                    (case try (Graph.get_node access_G) (nickname_of_thm th) of
wneuper@59324
  1455
                      SOME (Isar_Proof, _, deps) => ~100 * length deps
wneuper@59324
  1456
                    | SOME (Automatic_Proof, _, _) => 2 * max_isar
wneuper@59324
  1457
                    | SOME (Isar_Proof_wegen_Prover_Flop, _, _) => max_isar
wneuper@59324
  1458
                    | NONE => 0)
blanchet@58347
  1459
wneuper@59324
  1460
                  val old_facts = facts
wneuper@59324
  1461
                    |> filter is_in_access_G
wneuper@59324
  1462
                    |> map (`(priority_of o snd))
wneuper@59324
  1463
                    |> sort (int_ord o apply2 fst)
wneuper@59324
  1464
                    |> map snd
wneuper@59324
  1465
                  val ((relearns, flops), (num_nontrivial, _, _)) =
wneuper@59324
  1466
                    (([], []), (num_nontrivial, next_commit_time (), false))
wneuper@59324
  1467
                    |> fold relearn_old_fact old_facts
wneuper@59324
  1468
                in
wneuper@59324
  1469
                  commit true [] relearns flops; num_nontrivial
wneuper@59324
  1470
                end
wneuper@59324
  1471
          in
wneuper@59324
  1472
            if verbose orelse auto_level < 2 then
wneuper@59324
  1473
              "Learned " ^ string_of_int num_new_facts ^ " fact" ^ plural_s num_new_facts ^
wneuper@59324
  1474
              " and " ^ string_of_int num_nontrivial ^ " nontrivial " ^
wneuper@59324
  1475
              (if run_prover then "automatic and " else "") ^ "Isar proof" ^
wneuper@59324
  1476
              plural_s num_nontrivial ^
wneuper@59324
  1477
              (if verbose then " in " ^ string_of_time (Timer.checkRealTimer timer) else "")
wneuper@59324
  1478
            else
wneuper@59324
  1479
              ""
wneuper@59324
  1480
          end
wneuper@59324
  1481
      end)
blanchet@49323
  1482
  end
blanchet@49319
  1483
blanchet@55575
  1484
fun mash_learn ctxt (params as {provers, timeout, ...}) fact_override chained run_prover =
blanchet@49331
  1485
  let
blanchet@49411
  1486
    val css = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
blanchet@49410
  1487
    val ctxt = ctxt |> Config.put instantiate_inducts false
wneuper@59180
  1488
    val facts =
walther@59606
  1489
      nearly_all_facts ctxt false fact_override Keyword.empty_keywords css chained [] \<^prop>\<open>True\<close>
wneuper@59324
  1490
      |> sort (crude_thm_ord ctxt o apply2 snd o swap)
blanchet@49419
  1491
    val num_facts = length facts
blanchet@49419
  1492
    val prover = hd provers
blanchet@58348
  1493
blanchet@51499
  1494
    fun learn auto_level run_prover =
blanchet@58773
  1495
      mash_learn_facts ctxt params prover auto_level run_prover one_year facts
wneuper@59180
  1496
      |> writeln
blanchet@49331
  1497
  in
blanchet@51499
  1498
    if run_prover then
wneuper@59180
  1499
      (writeln ("MaShing through " ^ string_of_int num_facts ^ " fact" ^
blanchet@58359
  1500
         plural_s num_facts ^ " for automatic proofs (" ^ quote prover ^ " timeout: " ^
blanchet@58359
  1501
         string_of_time timeout ^ ").\n\nCollecting Isar proofs first...");
blanchet@51355
  1502
       learn 1 false;
wneuper@59324
  1503
       writeln "Now collecting automatic proofs\n\
wneuper@59324
  1504
         \This may take several hours; you can safely stop the learning process at any point";
blanchet@51355
  1505
       learn 0 true)
blanchet@51355
  1506
    else
wneuper@59180
  1507
      (writeln ("MaShing through " ^ string_of_int num_facts ^ " fact" ^
blanchet@58348
  1508
         plural_s num_facts ^ " for Isar proofs...");
blanchet@51355
  1509
       learn 0 false)
blanchet@49331
  1510
  end
blanchet@49264
  1511
wneuper@59324
  1512
fun mash_can_suggest_facts ctxt =
wneuper@59324
  1513
  (case get_state ctxt of
wneuper@59324
  1514
    NONE => false
wneuper@59324
  1515
  | SOME {access_G, ...} => not (Graph.is_empty access_G))
wneuper@59324
  1516
wneuper@59324
  1517
fun mash_can_suggest_facts_fast ctxt =
wneuper@59324
  1518
  (case peek_state ctxt of
wneuper@59324
  1519
    NONE => false
wneuper@59324
  1520
  | SOME (_, {access_G, ...}) => not (Graph.is_empty access_G))
blanchet@51326
  1521
blanchet@58616
  1522
(* Generate more suggestions than requested, because some might be thrown out later for various
blanchet@58616
  1523
   reasons (e.g., duplicates). *)
walther@59606
  1524
fun generous_max_suggestions max_facts = 2 * max_facts + 25 (* FUDGE *)
blanchet@51398
  1525
wneuper@59324
  1526
val mepo_weight = 0.5 (* FUDGE *)
wneuper@59324
  1527
val mash_weight = 0.5 (* FUDGE *)
blanchet@51829
  1528
wneuper@59324
  1529
val max_facts_to_learn_before_query = 100 (* FUDGE *)
blanchet@54289
  1530
wneuper@59324
  1531
(* The threshold should be large enough so that MaSh does not get activated for Auto Sledgehammer. *)
wneuper@59324
  1532
val min_secs_for_learning = 10
blanchet@49333
  1533
blanchet@58773
  1534
fun relevant_facts ctxt (params as {verbose, learn, fact_filter, timeout, ...}) prover
blanchet@58747
  1535
    max_facts ({add, only, ...} : fact_override) hyp_ts concl_t facts =
blanchet@49329
  1536
  if not (subset (op =) (the_list fact_filter, fact_filters)) then
wneuper@59324
  1537
    error ("Unknown fact filter: " ^ quote (the fact_filter))
blanchet@49329
  1538
  else if only then
blanchet@58492
  1539
    [("", map fact_of_raw_fact facts)]
blanchet@49336
  1540
  else if max_facts <= 0 orelse null facts then
blanchet@52192
  1541
    [("", [])]
blanchet@49303
  1542
  else
blanchet@49303
  1543
    let
wneuper@59324
  1544
      val thy_name = Context.theory_name (Proof_Context.theory_of ctxt)
blanchet@58749
  1545
blanchet@58897
  1546
      fun maybe_launch_thread exact min_num_facts_to_learn =
wneuper@59180
  1547
        if not (Async_Manager_Legacy.has_running_threads MaShN) andalso
blanchet@56158
  1548
           Time.toSeconds timeout >= min_secs_for_learning then
blanchet@56158
  1549
          let val timeout = time_mult learn_timeout_slack timeout in
blanchet@58747
  1550
            (if verbose then
wneuper@59180
  1551
               writeln ("Started MaShing through " ^
wneuper@59324
  1552
                 (if exact then "" else "up to ") ^ string_of_int min_num_facts_to_learn ^
wneuper@59324
  1553
                 " fact" ^ plural_s min_num_facts_to_learn ^ " in the background")
blanchet@58747
  1554
             else
blanchet@58747
  1555
               ());
blanchet@56158
  1556
            launch_thread timeout
blanchet@58773
  1557
              (fn () => (true, mash_learn_facts ctxt params prover 2 false timeout facts))
blanchet@49334
  1558
          end
blanchet@49333
  1559
        else
blanchet@49333
  1560
          ()
blanchet@58360
  1561
wneuper@59324
  1562
      val mash_enabled = is_mash_enabled ()
wneuper@59324
  1563
      val mash_fast = mash_can_suggest_facts_fast ctxt
wneuper@59324
  1564
blanchet@58897
  1565
      fun please_learn () =
wneuper@59324
  1566
        if mash_fast then
wneuper@59324
  1567
          (case get_state ctxt of
wneuper@59324
  1568
            NONE => maybe_launch_thread false (length facts)
wneuper@59324
  1569
          | SOME {access_G, xtabs = ((num_facts0, _), _), ...} =>
wneuper@59324
  1570
            let
wneuper@59324
  1571
              val is_in_access_G = is_fact_in_graph access_G o snd
wneuper@59324
  1572
              val min_num_facts_to_learn = length facts - num_facts0
wneuper@59324
  1573
            in
wneuper@59324
  1574
              if min_num_facts_to_learn <= max_facts_to_learn_before_query then
wneuper@59324
  1575
                (case length (filter_out is_in_access_G facts) of
wneuper@59324
  1576
                  0 => ()
wneuper@59324
  1577
                | num_facts_to_learn =>
wneuper@59324
  1578
                  if num_facts_to_learn <= max_facts_to_learn_before_query then
wneuper@59324
  1579
                    mash_learn_facts ctxt params prover 2 false timeout facts
wneuper@59324
  1580
                    |> (fn "" => () | s => writeln (MaShN ^ ": " ^ s))
wneuper@59324
  1581
                  else
wneuper@59324
  1582
                    maybe_launch_thread true num_facts_to_learn)
blanchet@58897
  1583
              else
wneuper@59324
  1584
                maybe_launch_thread false min_num_facts_to_learn
wneuper@59324
  1585
            end)
wneuper@59324
  1586
        else
wneuper@59324
  1587
          maybe_launch_thread false (length facts)
blanchet@58897
  1588
blanchet@58897
  1589
      val _ =
blanchet@58897
  1590
        if learn andalso mash_enabled andalso fact_filter <> SOME mepoN then please_learn () else ()
blanchet@58360
  1591
blanchet@58773
  1592
      val effective_fact_filter =
blanchet@56628
  1593
        (case fact_filter of
blanchet@58773
  1594
          SOME ff => ff
wneuper@59324
  1595
        | NONE => if mash_enabled andalso mash_fast then meshN else mepoN)
blanchet@55595
  1596
blanchet@55595
  1597
      val unique_facts = drop_duplicate_facts facts
blanchet@49303
  1598
      val add_ths = Attrib.eval_thms ctxt add
blanchet@55595
  1599
blanchet@52186
  1600
      fun in_add (_, th) = member Thm.eq_thm_prop add_ths th
blanchet@58360
  1601
blanchet@52185
  1602
      fun add_and_take accepts =
blanchet@52185
  1603
        (case add_ths of
blanchet@52185
  1604
           [] => accepts
blanchet@58492
  1605
         | _ =>
blanchet@58492
  1606
           (unique_facts |> filter in_add |> map fact_of_raw_fact) @ (accepts |> filter_out in_add))
blanchet@49308
  1607
        |> take max_facts
blanchet@58360
  1608
blanchet@49421
  1609
      fun mepo () =
blanchet@55595
  1610
        (mepo_suggested_facts ctxt params max_facts NONE hyp_ts concl_t unique_facts
blanchet@55543
  1611
         |> weight_facts_steeply, [])
blanchet@58360
  1612
blanchet@49329
  1613
      fun mash () =
wneuper@59324
  1614
        mash_suggested_facts ctxt thy_name params (generous_max_suggestions max_facts) hyp_ts
wneuper@59324
  1615
          concl_t facts
blanchet@54277
  1616
        |>> weight_facts_steeply
blanchet@58360
  1617
blanchet@49329
  1618
      val mess =
blanchet@52185
  1619
        (* the order is important for the "case" expression below *)
blanchet@55543
  1620
        [] |> effective_fact_filter <> mepoN ? cons (mash_weight, mash)
blanchet@55543
  1621
           |> effective_fact_filter <> mashN ? cons (mepo_weight, mepo)
blanchet@55543
  1622
           |> Par_List.map (apsnd (fn f => f ()))
wneuper@59324
  1623
      val mesh =
wneuper@59324
  1624
        mesh_facts (fact_distinct (op aconv)) (eq_snd (gen_eq_thm ctxt)) max_facts mess
wneuper@59324
  1625
        |> add_and_take
blanchet@49303
  1626
    in
blanchet@56628
  1627
      (case (fact_filter, mess) of
blanchet@52206
  1628
        (NONE, [(_, (mepo, _)), (_, (mash, _))]) =>
blanchet@59000
  1629
        [(meshN, mesh),
blanchet@59000
  1630
         (mepoN, mepo |> map fst |> add_and_take),
blanchet@52192
  1631
         (mashN, mash |> map fst |> add_and_take)]
blanchet@56628
  1632
      | _ => [(effective_fact_filter, mesh)])
blanchet@49303
  1633
    end
blanchet@49303
  1634
blanchet@49263
  1635
end;