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