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