src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 49345 192444b53e86
parent 49344 5781c4620245
child 49347 271a4a6af734
permissions -rw-r--r--
speed up MaSh queries
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3 
     4 Sledgehammer's machine-learning-based relevance filter (MaSh).
     5 *)
     6 
     7 signature SLEDGEHAMMER_FILTER_MASH =
     8 sig
     9   type status = ATP_Problem_Generate.status
    10   type stature = ATP_Problem_Generate.stature
    11   type fact = Sledgehammer_Fact.fact
    12   type fact_override = Sledgehammer_Fact.fact_override
    13   type params = Sledgehammer_Provers.params
    14   type relevance_fudge = Sledgehammer_Provers.relevance_fudge
    15   type prover_result = Sledgehammer_Provers.prover_result
    16 
    17   val trace : bool Config.T
    18   val MaShN : string
    19   val meshN : string
    20   val iterN : string
    21   val mashN : string
    22   val fact_filters : string list
    23   val escape_meta : string -> string
    24   val escape_metas : string list -> string
    25   val unescape_meta : string -> string
    26   val unescape_metas : string -> string list
    27   val extract_query : string -> string * string list
    28   val suggested_facts : string list -> ('a * thm) list -> ('a * thm) list
    29   val mesh_facts :
    30     int -> (('a * thm) list * ('a * thm) list) list -> ('a * thm) list
    31   val is_likely_tautology_or_too_meta : thm -> bool
    32   val theory_ord : theory * theory -> order
    33   val thm_ord : thm * thm -> order
    34   val features_of :
    35     Proof.context -> string -> theory -> status -> term list -> string list
    36   val isabelle_dependencies_of : unit Symtab.table -> thm -> string list
    37   val goal_of_thm : theory -> thm -> thm
    38   val run_prover_for_mash :
    39     Proof.context -> params -> string -> fact list -> thm -> prover_result
    40   val mash_RESET : Proof.context -> unit
    41   val mash_INIT :
    42     Proof.context -> bool
    43     -> (string * string list * string list * string list) list -> unit
    44   val mash_ADD :
    45     Proof.context -> bool
    46     -> (string * string list * string list * string list) list -> unit
    47   val mash_QUERY :
    48     Proof.context -> bool -> int -> string list * string list -> string list
    49   val mash_reset : Proof.context -> unit
    50   val mash_could_suggest_facts : unit -> bool
    51   val mash_can_suggest_facts : Proof.context -> bool
    52   val mash_suggest_facts :
    53     Proof.context -> params -> string -> int -> term list -> term
    54     -> ('a * thm) list -> ('a * thm) list * ('a * thm) list
    55   val mash_learn_thy :
    56     Proof.context -> params -> theory -> Time.time -> fact list -> string
    57   val mash_learn_proof :
    58     Proof.context -> params -> term -> ('a * thm) list -> thm list -> unit
    59   val relevant_facts :
    60     Proof.context -> params -> string -> int -> fact_override -> term list
    61     -> term -> fact list -> fact list
    62   val kill_learners : unit -> unit
    63   val running_learners : unit -> unit
    64 end;
    65 
    66 structure Sledgehammer_Filter_MaSh : SLEDGEHAMMER_FILTER_MASH =
    67 struct
    68 
    69 open ATP_Util
    70 open ATP_Problem_Generate
    71 open Sledgehammer_Util
    72 open Sledgehammer_Fact
    73 open Sledgehammer_Filter_Iter
    74 open Sledgehammer_Provers
    75 open Sledgehammer_Minimize
    76 
    77 val trace =
    78   Attrib.setup_config_bool @{binding sledgehammer_filter_mash_trace} (K false)
    79 fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
    80 
    81 val MaShN = "MaSh"
    82 
    83 val meshN = "mesh"
    84 val iterN = "iter"
    85 val mashN = "mash"
    86 
    87 val fact_filters = [meshN, iterN, mashN]
    88 
    89 fun mash_home () = getenv "MASH_HOME"
    90 fun mash_state_dir () =
    91   getenv "ISABELLE_HOME_USER" ^ "/mash"
    92   |> tap (Isabelle_System.mkdir o Path.explode)
    93 fun mash_state_path () = mash_state_dir () ^ "/state" |> Path.explode
    94 
    95 
    96 (*** Isabelle helpers ***)
    97 
    98 fun meta_char c =
    99   if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
   100      c = #")" orelse c = #"," then
   101     String.str c
   102   else
   103     (* fixed width, in case more digits follow *)
   104     "\\" ^ stringN_of_int 3 (Char.ord c)
   105 
   106 fun unmeta_chars accum [] = String.implode (rev accum)
   107   | unmeta_chars accum (#"\\" :: d1 :: d2 :: d3 :: cs) =
   108     (case Int.fromString (String.implode [d1, d2, d3]) of
   109        SOME n => unmeta_chars (Char.chr n :: accum) cs
   110      | NONE => "" (* error *))
   111   | unmeta_chars _ (#"\\" :: _) = "" (* error *)
   112   | unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs
   113 
   114 val escape_meta = String.translate meta_char
   115 val escape_metas = map escape_meta #> space_implode " "
   116 val unescape_meta = String.explode #> unmeta_chars []
   117 val unescape_metas =
   118   space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta
   119 
   120 fun extract_query line =
   121   case space_explode ":" line of
   122     [goal_name, suggs] => (unescape_meta goal_name, unescape_metas suggs)
   123   | _ => ("", [])
   124 
   125 fun suggested_facts suggs facts =
   126   let
   127     fun add_fact (fact as (_, th)) = Symtab.default (Thm.get_name_hint th, fact)
   128     val tab = Symtab.empty |> fold add_fact facts
   129   in map_filter (Symtab.lookup tab) suggs end
   130 
   131 (* Ad hoc score function roughly based on Blanchette's Ringberg 2011 data. *)
   132 fun score x = Math.pow (1.5, 15.5 - 0.05 * Real.fromInt x) + 15.0
   133 
   134 fun sum_sq_avg [] = 0
   135   | sum_sq_avg xs =
   136     Real.ceil (100000.0 * fold (curry (op +) o score) xs 0.0) div length xs
   137 
   138 fun mesh_facts max_facts [(selected, unknown)] =
   139     take max_facts selected @ take (max_facts - length selected) unknown
   140   | mesh_facts max_facts mess =
   141     let
   142       val mess = mess |> map (apfst (`length))
   143       val fact_eq = Thm.eq_thm o pairself snd
   144       fun score_in fact ((sel_len, sels), unks) =
   145         case find_index (curry fact_eq fact) sels of
   146           ~1 => (case find_index (curry fact_eq fact) unks of
   147                    ~1 => SOME sel_len
   148                  | _ => NONE)
   149         | j => SOME j
   150       fun score_of fact = mess |> map_filter (score_in fact) |> sum_sq_avg
   151       val facts = fold (union fact_eq o take max_facts o snd o fst) mess []
   152     in
   153       facts |> map (`score_of) |> sort (int_ord o swap o pairself fst)
   154             |> map snd |> take max_facts
   155     end
   156 
   157 val thy_feature_prefix = "y_"
   158 
   159 val thy_feature_name_of = prefix thy_feature_prefix
   160 val const_name_of = prefix const_prefix
   161 val type_name_of = prefix type_const_prefix
   162 val class_name_of = prefix class_prefix
   163 
   164 fun is_likely_tautology_or_too_meta th =
   165   let
   166     val is_boring_const = member (op =) atp_widely_irrelevant_consts
   167     fun is_boring_bool t =
   168       not (exists_Const (not o is_boring_const o fst) t) orelse
   169       exists_type (exists_subtype (curry (op =) @{typ prop})) t
   170     fun is_boring_prop (@{const Trueprop} $ t) = is_boring_bool t
   171       | is_boring_prop (@{const "==>"} $ t $ u) =
   172         is_boring_prop t andalso is_boring_prop u
   173       | is_boring_prop (Const (@{const_name all}, _) $ (Abs (_, _, t)) $ u) =
   174         is_boring_prop t andalso is_boring_prop u
   175       | is_boring_prop (Const (@{const_name "=="}, _) $ t $ u) =
   176         is_boring_bool t andalso is_boring_bool u
   177       | is_boring_prop _ = true
   178   in
   179     is_boring_prop (prop_of th) andalso not (Thm.eq_thm_prop (@{thm ext}, th))
   180   end
   181 
   182 fun theory_ord p =
   183   if Theory.eq_thy p then
   184     EQUAL
   185   else if Theory.subthy p then
   186     LESS
   187   else if Theory.subthy (swap p) then
   188     GREATER
   189   else case int_ord (pairself (length o Theory.ancestors_of) p) of
   190     EQUAL => string_ord (pairself Context.theory_name p)
   191   | order => order
   192 
   193 val thm_ord = theory_ord o pairself theory_of_thm
   194 
   195 val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
   196 
   197 fun interesting_terms_types_and_classes ctxt prover term_max_depth
   198                                         type_max_depth ts =
   199   let
   200     fun is_bad_const (x as (s, _)) args =
   201       member (op =) atp_logical_consts s orelse
   202       fst (is_built_in_const_for_prover ctxt prover x args)
   203     fun add_classes @{sort type} = I
   204       | add_classes S = union (op =) (map class_name_of S)
   205     fun do_add_type (Type (s, Ts)) =
   206         (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
   207         #> fold do_add_type Ts
   208       | do_add_type (TFree (_, S)) = add_classes S
   209       | do_add_type (TVar (_, S)) = add_classes S
   210     fun add_type T = type_max_depth >= 0 ? do_add_type T
   211     fun mk_app s args =
   212       if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
   213       else s
   214     fun patternify ~1 _ = ""
   215       | patternify depth t =
   216         case strip_comb t of
   217           (Const (s, _), args) =>
   218           mk_app (const_name_of s) (map (patternify (depth - 1)) args)
   219         | _ => ""
   220     fun add_term_patterns ~1 _ = I
   221       | add_term_patterns depth t =
   222         insert (op =) (patternify depth t)
   223         #> add_term_patterns (depth - 1) t
   224     val add_term = add_term_patterns term_max_depth
   225     fun add_patterns t =
   226       let val (head, args) = strip_comb t in
   227         (case head of
   228            Const (x as (_, T)) =>
   229            not (is_bad_const x args) ? (add_term t #> add_type T)
   230          | Free (_, T) => add_type T
   231          | Var (_, T) => add_type T
   232          | Abs (_, T, body) => add_type T #> add_patterns body
   233          | _ => I)
   234         #> fold add_patterns args
   235       end
   236   in [] |> fold add_patterns ts end
   237 
   238 fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
   239 
   240 val term_max_depth = 1
   241 val type_max_depth = 1
   242 
   243 (* TODO: Generate type classes for types? *)
   244 fun features_of ctxt prover thy status ts =
   245   thy_feature_name_of (Context.theory_name thy) ::
   246   interesting_terms_types_and_classes ctxt prover term_max_depth type_max_depth
   247                                       ts
   248   |> exists (not o is_lambda_free) ts ? cons "lambdas"
   249   |> exists (exists_Const is_exists) ts ? cons "skolems"
   250   |> (case status of
   251         General => I
   252       | Induction => cons "induction"
   253       | Intro => cons "intro"
   254       | Inductive => cons "inductive"
   255       | Elim => cons "elim"
   256       | Simp => cons "simp"
   257       | Def => cons "def")
   258 
   259 fun isabelle_dependencies_of all_facts = thms_in_proof (SOME all_facts)
   260 
   261 val freezeT = Type.legacy_freeze_type
   262 
   263 fun freeze (t $ u) = freeze t $ freeze u
   264   | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
   265   | freeze (Var ((s, _), T)) = Free (s, freezeT T)
   266   | freeze (Const (s, T)) = Const (s, freezeT T)
   267   | freeze (Free (s, T)) = Free (s, freezeT T)
   268   | freeze t = t
   269 
   270 fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
   271 
   272 fun run_prover_for_mash ctxt params prover facts goal =
   273   let
   274     val problem =
   275       {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
   276        facts = facts |> map (apfst (apfst (fn name => name ())))
   277                      |> map Untranslated_Fact}
   278     val prover = get_minimizing_prover ctxt Normal (K ()) prover
   279   in prover params (K (K (K ""))) problem end
   280 
   281 
   282 (*** Low-level communication with MaSh ***)
   283 
   284 fun write_file (xs, f) file =
   285   let val path = Path.explode file in
   286     File.write path "";
   287     xs |> chunk_list 500
   288        |> List.app (File.append path o space_implode "" o map f)
   289   end
   290 
   291 fun mash_info overlord =
   292   if overlord then (getenv "ISABELLE_HOME_USER", "")
   293   else (getenv "ISABELLE_TMP", serial_string ())
   294 
   295 fun run_mash ctxt (temp_dir, serial) core =
   296   let
   297     val log_file = temp_dir ^ "/mash_log" ^ serial
   298     val err_file = temp_dir ^ "/mash_err" ^ serial
   299     val command =
   300       mash_home () ^ "/mash.py --quiet --outputDir " ^ mash_state_dir () ^
   301       " --log " ^ log_file ^ " " ^ core ^ " 2>&1 > " ^ err_file
   302   in
   303     trace_msg ctxt (fn () => "Running " ^ command);
   304     write_file ([], K "") log_file;
   305     write_file ([], K "") err_file;
   306     Isabelle_System.bash command; ()
   307   end
   308 
   309 fun run_mash_init ctxt overlord write_access write_feats write_deps =
   310   let
   311     val info as (temp_dir, serial) = mash_info overlord
   312     val in_dir = temp_dir ^ "/mash_init" ^ serial
   313                  |> tap (Isabelle_System.mkdir o Path.explode)
   314   in
   315     write_file write_access (in_dir ^ "/mash_accessibility");
   316     write_file write_feats (in_dir ^ "/mash_features");
   317     write_file write_deps (in_dir ^ "/mash_dependencies");
   318     run_mash ctxt info ("--init --inputDir " ^ in_dir)
   319   end
   320 
   321 fun run_mash_commands ctxt overlord save max_suggs write_cmds read_suggs =
   322   let
   323     val info as (temp_dir, serial) = mash_info overlord
   324     val sugg_file = temp_dir ^ "/mash_suggs" ^ serial
   325     val cmd_file = temp_dir ^ "/mash_commands" ^ serial
   326   in
   327     write_file ([], K "") sugg_file;
   328     write_file write_cmds cmd_file;
   329     run_mash ctxt info
   330              ("--inputFile " ^ cmd_file ^ " --predictions " ^ sugg_file ^
   331               " --numberOfPredictions " ^ string_of_int max_suggs ^
   332               (if save then " --saveModel" else ""));
   333     read_suggs (fn () => File.read_lines (Path.explode sugg_file))
   334   end
   335 
   336 fun str_of_update (name, parents, feats, deps) =
   337   "! " ^ escape_meta name ^ ": " ^ escape_metas parents ^ "; " ^
   338   escape_metas feats ^ "; " ^ escape_metas deps ^ "\n"
   339 
   340 fun str_of_query (parents, feats) =
   341   "? " ^ escape_metas parents ^ "; " ^ escape_metas feats
   342 
   343 fun init_str_of_update get (upd as (name, _, _, _)) =
   344   escape_meta name ^ ": " ^ escape_metas (get upd) ^ "\n"
   345 
   346 fun mash_RESET ctxt =
   347   let val path = mash_state_dir () |> Path.explode in
   348     trace_msg ctxt (K "MaSh RESET");
   349     File.fold_dir (fn file => fn () =>
   350                       File.rm (Path.append path (Path.basic file)))
   351                   path ()
   352   end
   353 
   354 fun mash_INIT ctxt _ [] = mash_RESET ctxt
   355   | mash_INIT ctxt overlord upds =
   356     (trace_msg ctxt (fn () => "MaSh INIT " ^
   357          elide_string 1000 (space_implode " " (map #1 upds)));
   358      run_mash_init ctxt overlord (upds, init_str_of_update #2)
   359                    (upds, init_str_of_update #3) (upds, init_str_of_update #4))
   360 
   361 fun mash_ADD _ _ [] = ()
   362   | mash_ADD ctxt overlord upds =
   363     (trace_msg ctxt (fn () => "MaSh ADD " ^
   364          elide_string 1000 (space_implode " " (map #1 upds)));
   365      run_mash_commands ctxt overlord true 0 (upds, str_of_update) (K ()))
   366 
   367 fun mash_QUERY ctxt overlord max_suggs (query as (_, feats)) =
   368   (trace_msg ctxt (fn () => "MaSh QUERY " ^ space_implode " " feats);
   369    run_mash_commands ctxt overlord false max_suggs
   370        ([query], str_of_query)
   371        (fn suggs => snd (extract_query (List.last (suggs ()))))
   372    handle List.Empty => [])
   373 
   374 
   375 (*** High-level communication with MaSh ***)
   376 
   377 fun try_graph ctxt when def f =
   378   f ()
   379   handle Graph.CYCLES (cycle :: _) =>
   380          (trace_msg ctxt (fn () =>
   381               "Cycle involving " ^ commas cycle ^ " when " ^ when); def)
   382        | Graph.UNDEF name =>
   383          (trace_msg ctxt (fn () =>
   384               "Unknown fact " ^ quote name ^ " when " ^ when); def)
   385 
   386 type mash_state =
   387   {thys : bool Symtab.table,
   388    fact_graph : unit Graph.T}
   389 
   390 val empty_state = {thys = Symtab.empty, fact_graph = Graph.empty}
   391 
   392 local
   393 
   394 fun mash_load _ (state as (true, _)) = state
   395   | mash_load ctxt _ =
   396     let val path = mash_state_path () in
   397       (true,
   398        case try File.read_lines path of
   399          SOME (comp_thys :: incomp_thys :: fact_lines) =>
   400          let
   401            fun add_thy comp thy = Symtab.update (thy, comp)
   402            fun add_edge_to name parent =
   403              Graph.default_node (parent, ())
   404              #> Graph.add_edge (parent, name)
   405            fun add_fact_line line =
   406              case extract_query line of
   407                ("", _) => I (* shouldn't happen *)
   408              | (name, parents) =>
   409                Graph.default_node (name, ())
   410                #> fold (add_edge_to name) parents
   411            val thys =
   412              Symtab.empty |> fold (add_thy true) (unescape_metas comp_thys)
   413                           |> fold (add_thy false) (unescape_metas incomp_thys)
   414            val fact_graph =
   415              try_graph ctxt "loading state" Graph.empty (fn () =>
   416                  Graph.empty |> fold add_fact_line fact_lines)
   417          in {thys = thys, fact_graph = fact_graph} end
   418        | _ => empty_state)
   419     end
   420 
   421 fun mash_save ({thys, fact_graph, ...} : mash_state) =
   422   let
   423     val path = mash_state_path ()
   424     val thys = Symtab.dest thys
   425     val line_for_thys = escape_metas o AList.find (op =) thys
   426     fun fact_line_for name parents =
   427       escape_meta name ^ ": " ^ escape_metas parents
   428     val append_fact = File.append path o suffix "\n" oo fact_line_for
   429   in
   430     File.write path (line_for_thys true ^ "\n" ^ line_for_thys false ^ "\n");
   431     Graph.fold (fn (name, ((), (parents, _))) => fn () =>
   432                    append_fact name (Graph.Keys.dest parents))
   433         fact_graph ()
   434   end
   435 
   436 val global_state =
   437   Synchronized.var "Sledgehammer_Filter_MaSh.global_state" (false, empty_state)
   438 
   439 in
   440 
   441 fun mash_map ctxt f =
   442   Synchronized.change global_state (mash_load ctxt ##> (f #> tap mash_save))
   443 
   444 fun mash_get ctxt =
   445   Synchronized.change_result global_state (mash_load ctxt #> `snd)
   446 
   447 fun mash_reset ctxt =
   448   Synchronized.change global_state (fn _ =>
   449       (mash_RESET ctxt; File.write (mash_state_path ()) "";
   450        (true, empty_state)))
   451 
   452 end
   453 
   454 fun mash_could_suggest_facts () = mash_home () <> ""
   455 fun mash_can_suggest_facts ctxt =
   456   not (Graph.is_empty (#fact_graph (mash_get ctxt)))
   457 
   458 fun parents_wrt_facts ctxt facts fact_graph =
   459   let
   460     val facts = [] |> fold (cons o Thm.get_name_hint o snd) facts
   461     val tab = Symtab.empty |> fold (fn name => Symtab.update (name, ())) facts
   462     val maxs = Graph.maximals fact_graph
   463   in
   464     if forall (Symtab.defined tab) maxs then
   465       maxs
   466     else
   467       let
   468         val graph_facts = Symtab.make (map (rpair ()) (Graph.keys fact_graph))
   469         val ancestors =
   470           try_graph ctxt "when computing ancestor facts" [] (fn () =>
   471               facts |> filter (Symtab.defined graph_facts)
   472                     |> Graph.all_preds fact_graph)
   473         val ancestors =
   474           Symtab.empty |> fold (fn name => Symtab.update (name, ())) ancestors
   475       in
   476         try_graph ctxt "when computing parent facts" [] (fn () =>
   477             fact_graph |> Graph.restrict (Symtab.defined ancestors)
   478                        |> Graph.maximals)
   479       end
   480   end
   481 
   482 (* Generate more suggestions than requested, because some might be thrown out
   483    later for various reasons and "meshing" gives better results with some
   484    slack. *)
   485 fun max_suggs_of max_facts = max_facts + Int.min (200, max_facts)
   486 
   487 fun is_fact_in_graph fact_graph (_, th) =
   488   can (Graph.get_node fact_graph) (Thm.get_name_hint th)
   489 
   490 fun mash_suggest_facts ctxt ({overlord, ...} : params) prover max_facts hyp_ts
   491                        concl_t facts =
   492   let
   493 val timer = Timer.startRealTimer ()
   494     val thy = Proof_Context.theory_of ctxt
   495 val _ = tracing ("A1: " ^ PolyML.makestring (Timer.checkRealTimer timer)) (*###*)
   496     val fact_graph = #fact_graph (mash_get ctxt)
   497 val _ = tracing ("A2: " ^ PolyML.makestring (Timer.checkRealTimer timer)) (*###*)
   498     val parents = parents_wrt_facts ctxt facts fact_graph
   499 val _ = tracing ("A3: " ^ PolyML.makestring (Timer.checkRealTimer timer)) (*###*)
   500     val feats = features_of ctxt prover thy General (concl_t :: hyp_ts)
   501 val _ = tracing ("A4: " ^ PolyML.makestring (Timer.checkRealTimer timer)) (*###*)
   502     val suggs =
   503       if Graph.is_empty fact_graph then []
   504       else mash_QUERY ctxt overlord (max_suggs_of max_facts) (parents, feats)
   505 val _ = tracing ("A5: " ^ PolyML.makestring (Timer.checkRealTimer timer)) (*###*)
   506     val selected = facts |> suggested_facts suggs
   507 val _ = tracing ("A6: " ^ PolyML.makestring (Timer.checkRealTimer timer)) (*###*)
   508     val unknown = facts |> filter_out (is_fact_in_graph fact_graph)
   509 val _ = tracing ("END: " ^ PolyML.makestring (Timer.checkRealTimer timer)) (*###*)
   510   in (selected, unknown) end
   511 
   512 fun add_thys_for thy =
   513   let fun add comp thy = Symtab.update (Context.theory_name thy, comp) in
   514     add false thy #> fold (add true) (Theory.ancestors_of thy)
   515   end
   516 
   517 fun update_fact_graph ctxt (name, parents, feats, deps) (upds, graph) =
   518   let
   519     fun maybe_add_from from (accum as (parents, graph)) =
   520       try_graph ctxt "updating graph" accum (fn () =>
   521           (from :: parents, Graph.add_edge_acyclic (from, name) graph))
   522     val graph = graph |> Graph.default_node (name, ())
   523     val (parents, graph) = ([], graph) |> fold maybe_add_from parents
   524     val (deps, graph) = ([], graph) |> fold maybe_add_from deps
   525   in ((name, parents, feats, deps) :: upds, graph) end
   526 
   527 val pass1_learn_timeout_factor = 0.5
   528 
   529 (* The timeout is understood in a very slack fashion. *)
   530 fun mash_learn_thy ctxt ({provers, verbose, overlord, ...} : params) thy timeout
   531                    facts =
   532   let
   533     val timer = Timer.startRealTimer ()
   534     val prover = hd provers
   535     fun timed_out frac =
   536       Time.> (Timer.checkRealTimer timer, time_mult frac timeout)
   537     val {fact_graph, ...} = mash_get ctxt
   538     val new_facts =
   539       facts |> filter_out (is_fact_in_graph fact_graph)
   540             |> sort (thm_ord o pairself snd)
   541   in
   542     if null new_facts then
   543       ""
   544     else
   545       let
   546         val ths = facts |> map snd
   547         val all_names =
   548           ths |> filter_out is_likely_tautology_or_too_meta
   549               |> map (rpair () o Thm.get_name_hint)
   550               |> Symtab.make
   551         fun do_fact _ (accum as (_, true)) = accum
   552           | do_fact ((_, (_, status)), th) ((parents, upds), false) =
   553             let
   554               val name = Thm.get_name_hint th
   555               val feats = features_of ctxt prover thy status [prop_of th]
   556               val deps = isabelle_dependencies_of all_names th
   557               val upd = (name, parents, feats, deps)
   558             in (([name], upd :: upds), timed_out pass1_learn_timeout_factor) end
   559         val parents = parents_wrt_facts ctxt facts fact_graph
   560         val ((_, upds), _) =
   561           ((parents, []), false) |> fold do_fact new_facts |>> apsnd rev
   562         val n = length upds
   563         fun trans {thys, fact_graph} =
   564           let
   565             val mash_INIT_or_ADD =
   566               if Graph.is_empty fact_graph then mash_INIT else mash_ADD
   567             val (upds, fact_graph) =
   568               ([], fact_graph) |> fold (update_fact_graph ctxt) upds
   569           in
   570             (mash_INIT_or_ADD ctxt overlord (rev upds);
   571              {thys = thys |> add_thys_for thy,
   572               fact_graph = fact_graph})
   573           end
   574       in
   575         mash_map ctxt trans;
   576         if verbose then
   577           "Processed " ^ string_of_int n ^ " proof" ^ plural_s n ^
   578           (if verbose then
   579              " in " ^ string_from_time (Timer.checkRealTimer timer)
   580            else
   581              "") ^ "."
   582         else
   583           ""
   584       end
   585   end
   586 
   587 fun mash_learn_proof ctxt ({provers, overlord, ...} : params) t facts used_ths =
   588   let
   589     val thy = Proof_Context.theory_of ctxt
   590     val prover = hd provers
   591     val name = ATP_Util.timestamp () ^ " " ^ serial_string () (* fresh enough *)
   592     val feats = features_of ctxt prover thy General [t]
   593     val deps = used_ths |> map Thm.get_name_hint
   594   in
   595     mash_map ctxt (fn {thys, fact_graph} =>
   596         let
   597           val parents = parents_wrt_facts ctxt facts fact_graph
   598           val upds = [(name, parents, feats, deps)]
   599           val (upds, fact_graph) =
   600             ([], fact_graph) |> fold (update_fact_graph ctxt) upds
   601         in
   602           mash_ADD ctxt overlord upds;
   603           {thys = thys, fact_graph = fact_graph}
   604         end)
   605   end
   606 
   607 (* The threshold should be large enough so that MaSh doesn't kick in for Auto
   608    Sledgehammer and Try. *)
   609 val min_secs_for_learning = 15
   610 val learn_timeout_factor = 2.0
   611 
   612 fun relevant_facts ctxt (params as {learn, fact_filter, timeout, ...}) prover
   613         max_facts ({add, only, ...} : fact_override) hyp_ts concl_t facts =
   614   if not (subset (op =) (the_list fact_filter, fact_filters)) then
   615     error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".")
   616   else if only then
   617     facts
   618   else if max_facts <= 0 orelse null facts then
   619     []
   620   else
   621     let
   622       val thy = Proof_Context.theory_of ctxt
   623       fun maybe_learn () =
   624         if not learn orelse Async_Manager.has_running_threads MaShN then
   625           ()
   626         else if Time.toSeconds timeout >= min_secs_for_learning then
   627           let
   628             val soft_timeout = time_mult learn_timeout_factor timeout
   629             val hard_timeout = time_mult 4.0 soft_timeout
   630             val birth_time = Time.now ()
   631             val death_time = Time.+ (birth_time, hard_timeout)
   632             val desc = ("machine learner for Sledgehammer", "")
   633           in
   634             Async_Manager.launch MaShN birth_time death_time desc
   635                 (fn () =>
   636                     (true, mash_learn_thy ctxt params thy soft_timeout facts))
   637           end
   638         else
   639           ()
   640       val fact_filter =
   641         case fact_filter of
   642           SOME ff => (() |> ff <> iterN ? maybe_learn; ff)
   643         | NONE =>
   644           if mash_can_suggest_facts ctxt then (maybe_learn (); meshN)
   645           else if mash_could_suggest_facts () then (maybe_learn (); iterN)
   646           else iterN
   647       val add_ths = Attrib.eval_thms ctxt add
   648       fun prepend_facts ths accepts =
   649         ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
   650          (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
   651         |> take max_facts
   652       fun iter () =
   653         iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts
   654                                  concl_t facts
   655       fun mash () =
   656         mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts
   657       val mess =
   658         [] |> (if fact_filter <> mashN then cons (iter (), []) else I)
   659            |> (if fact_filter <> iterN then cons (mash ()) else I)
   660     in
   661       mesh_facts max_facts mess
   662       |> not (null add_ths) ? prepend_facts add_ths
   663     end
   664 
   665 fun kill_learners () = Async_Manager.kill_threads MaShN "learner"
   666 fun running_learners () = Async_Manager.running_threads MaShN "learner"
   667 
   668 end;