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