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