blanchet@49395: (* Title: HOL/Tools/Sledgehammer/sledgehammer_mash.ML blanchet@49263: Author: Jasmin Blanchette, TU Muenchen blanchet@49263: blanchet@49263: Sledgehammer's machine-learning-based relevance filter (MaSh). blanchet@49263: *) blanchet@49263: blanchet@49396: signature SLEDGEHAMMER_MASH = blanchet@49263: sig blanchet@49266: type status = ATP_Problem_Generate.status blanchet@49266: type stature = ATP_Problem_Generate.stature blanchet@49311: type fact = Sledgehammer_Fact.fact blanchet@49311: type fact_override = Sledgehammer_Fact.fact_override blanchet@49266: type params = Sledgehammer_Provers.params blanchet@49303: type relevance_fudge = Sledgehammer_Provers.relevance_fudge blanchet@49266: type prover_result = Sledgehammer_Provers.prover_result blanchet@49266: blanchet@49323: val trace : bool Config.T blanchet@49334: val MaShN : string blanchet@49394: val mepoN : string blanchet@49394: val mashN : string blanchet@49329: val meshN : string blanchet@49329: val fact_filters : string list blanchet@49318: val escape_meta : string -> string blanchet@49318: val escape_metas : string list -> string blanchet@49323: val unescape_meta : string -> string blanchet@49323: val unescape_metas : string -> string list blanchet@49326: val extract_query : string -> string * string list blanchet@49393: val nickname_of : thm -> string blanchet@49336: val suggested_facts : string list -> ('a * thm) list -> ('a * thm) list blanchet@49336: val mesh_facts : blanchet@49336: int -> (('a * thm) list * ('a * thm) list) list -> ('a * thm) list blanchet@49339: val is_likely_tautology_or_too_meta : thm -> bool blanchet@49266: val theory_ord : theory * theory -> order blanchet@49266: val thm_ord : thm * thm -> order blanchet@49333: val features_of : blanchet@49333: Proof.context -> string -> theory -> status -> term list -> string list blanchet@49331: val isabelle_dependencies_of : unit Symtab.table -> thm -> string list blanchet@49266: val goal_of_thm : theory -> thm -> thm blanchet@49336: val run_prover_for_mash : blanchet@49333: Proof.context -> params -> string -> fact list -> thm -> prover_result blanchet@49347: val mash_CLEAR : Proof.context -> unit blanchet@49331: val mash_INIT : blanchet@49331: Proof.context -> bool blanchet@49331: -> (string * string list * string list * string list) list -> unit blanchet@49323: val mash_ADD : blanchet@49331: Proof.context -> bool blanchet@49331: -> (string * string list * string list * string list) list -> unit blanchet@49331: val mash_QUERY : blanchet@49333: Proof.context -> bool -> int -> string list * string list -> string list blanchet@49347: val mash_unlearn : Proof.context -> unit blanchet@49333: val mash_could_suggest_facts : unit -> bool blanchet@49336: val mash_can_suggest_facts : Proof.context -> bool blanchet@49313: val mash_suggest_facts : blanchet@49336: Proof.context -> params -> string -> int -> term list -> term blanchet@49336: -> ('a * thm) list -> ('a * thm) list * ('a * thm) list blanchet@49334: val mash_learn_thy : blanchet@49334: Proof.context -> params -> theory -> Time.time -> fact list -> string blanchet@49331: val mash_learn_proof : blanchet@49336: Proof.context -> params -> term -> ('a * thm) list -> thm list -> unit blanchet@49303: val relevant_facts : blanchet@49307: Proof.context -> params -> string -> int -> fact_override -> term list blanchet@49311: -> term -> fact list -> fact list blanchet@49334: val kill_learners : unit -> unit blanchet@49334: val running_learners : unit -> unit blanchet@49263: end; blanchet@49263: blanchet@49396: structure Sledgehammer_MaSh : SLEDGEHAMMER_MASH = blanchet@49263: struct blanchet@49264: blanchet@49266: open ATP_Util blanchet@49266: open ATP_Problem_Generate blanchet@49266: open Sledgehammer_Util blanchet@49266: open Sledgehammer_Fact blanchet@49266: open Sledgehammer_Provers blanchet@49333: open Sledgehammer_Minimize blanchet@49396: open Sledgehammer_MePo blanchet@49266: blanchet@49323: val trace = blanchet@49395: Attrib.setup_config_bool @{binding sledgehammer_mash_trace} (K false) blanchet@49323: fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else () blanchet@49323: blanchet@49334: val MaShN = "MaSh" blanchet@49334: blanchet@49394: val mepoN = "mepo" blanchet@49394: val mashN = "mash" blanchet@49329: val meshN = "mesh" blanchet@49329: blanchet@49394: val fact_filters = [meshN, mepoN, mashN] blanchet@49329: blanchet@49329: fun mash_home () = getenv "MASH_HOME" blanchet@49329: fun mash_state_dir () = blanchet@49324: getenv "ISABELLE_HOME_USER" ^ "/mash" blanchet@49331: |> tap (Isabelle_System.mkdir o Path.explode) blanchet@49329: fun mash_state_path () = mash_state_dir () ^ "/state" |> Path.explode blanchet@49266: blanchet@49345: blanchet@49266: (*** Isabelle helpers ***) blanchet@49266: blanchet@49323: fun meta_char c = blanchet@49266: if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse blanchet@49266: c = #")" orelse c = #"," then blanchet@49266: String.str c blanchet@49266: else blanchet@49266: (* fixed width, in case more digits follow *) blanchet@49266: "\\" ^ stringN_of_int 3 (Char.ord c) blanchet@49266: blanchet@49323: fun unmeta_chars accum [] = String.implode (rev accum) blanchet@49323: | unmeta_chars accum (#"\\" :: d1 :: d2 :: d3 :: cs) = blanchet@49323: (case Int.fromString (String.implode [d1, d2, d3]) of blanchet@49323: SOME n => unmeta_chars (Char.chr n :: accum) cs blanchet@49323: | NONE => "" (* error *)) blanchet@49323: | unmeta_chars _ (#"\\" :: _) = "" (* error *) blanchet@49323: | unmeta_chars accum (c :: cs) = unmeta_chars (c :: accum) cs blanchet@49323: blanchet@49323: val escape_meta = String.translate meta_char blanchet@49318: val escape_metas = map escape_meta #> space_implode " " blanchet@49330: val unescape_meta = String.explode #> unmeta_chars [] blanchet@49330: val unescape_metas = blanchet@49330: space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta blanchet@49266: blanchet@49326: fun extract_query line = blanchet@49326: case space_explode ":" line of blanchet@49330: [goal_name, suggs] => (unescape_meta goal_name, unescape_metas suggs) blanchet@49327: | _ => ("", []) blanchet@49326: blanchet@49393: fun parent_of_local_thm th = blanchet@49393: let blanchet@49393: val thy = th |> Thm.theory_of_thm blanchet@49393: val facts = thy |> Global_Theory.facts_of blanchet@49393: val space = facts |> Facts.space_of blanchet@49393: fun id_of s = #id (Name_Space.the_entry space s) blanchet@49393: fun max_id (s', _) (s, id) = blanchet@49393: let val id' = id_of s' in if id > id' then (s, id) else (s', id') end blanchet@49393: in ("", ~1) |> Facts.fold_static max_id facts |> fst end blanchet@49393: blanchet@49393: val local_prefix = "local" ^ Long_Name.separator blanchet@49393: blanchet@49393: fun nickname_of th = blanchet@49393: let val hint = Thm.get_name_hint th in blanchet@49393: (* FIXME: There must be a better way to detect local facts. *) blanchet@49393: case try (unprefix local_prefix) hint of blanchet@49393: SOME suff => blanchet@49393: parent_of_local_thm th ^ Long_Name.separator ^ Long_Name.separator ^ suff blanchet@49393: | NONE => hint blanchet@49393: end blanchet@49393: blanchet@49345: fun suggested_facts suggs facts = blanchet@49345: let blanchet@49393: fun add_fact (fact as (_, th)) = Symtab.default (nickname_of th, fact) blanchet@49345: val tab = Symtab.empty |> fold add_fact facts blanchet@49345: in map_filter (Symtab.lookup tab) suggs end blanchet@49326: blanchet@49344: (* Ad hoc score function roughly based on Blanchette's Ringberg 2011 data. *) blanchet@49344: fun score x = Math.pow (1.5, 15.5 - 0.05 * Real.fromInt x) + 15.0 blanchet@49343: blanchet@49343: fun sum_sq_avg [] = 0 blanchet@49344: | sum_sq_avg xs = blanchet@49344: Real.ceil (100000.0 * fold (curry (op +) o score) xs 0.0) div length xs blanchet@49328: blanchet@49335: fun mesh_facts max_facts [(selected, unknown)] = blanchet@49335: take max_facts selected @ take (max_facts - length selected) unknown blanchet@49329: | mesh_facts max_facts mess = blanchet@49329: let blanchet@49335: val mess = mess |> map (apfst (`length)) blanchet@49329: val fact_eq = Thm.eq_thm o pairself snd blanchet@49335: fun score_in fact ((sel_len, sels), unks) = blanchet@49335: case find_index (curry fact_eq fact) sels of blanchet@49335: ~1 => (case find_index (curry fact_eq fact) unks of blanchet@49344: ~1 => SOME sel_len blanchet@49335: | _ => NONE) blanchet@49344: | j => SOME j blanchet@49343: fun score_of fact = mess |> map_filter (score_in fact) |> sum_sq_avg blanchet@49335: val facts = fold (union fact_eq o take max_facts o snd o fst) mess [] blanchet@49329: in blanchet@49343: facts |> map (`score_of) |> sort (int_ord o swap o pairself fst) blanchet@49343: |> map snd |> take max_facts blanchet@49329: end blanchet@49327: blanchet@49318: val thy_feature_prefix = "y_" blanchet@49266: blanchet@49318: val thy_feature_name_of = prefix thy_feature_prefix blanchet@49318: val const_name_of = prefix const_prefix blanchet@49318: val type_name_of = prefix type_const_prefix blanchet@49318: val class_name_of = prefix class_prefix blanchet@49266: blanchet@49339: fun is_likely_tautology_or_too_meta th = blanchet@49339: let blanchet@49339: val is_boring_const = member (op =) atp_widely_irrelevant_consts blanchet@49339: fun is_boring_bool t = blanchet@49339: not (exists_Const (not o is_boring_const o fst) t) orelse blanchet@49339: exists_type (exists_subtype (curry (op =) @{typ prop})) t blanchet@49339: fun is_boring_prop (@{const Trueprop} $ t) = is_boring_bool t blanchet@49339: | is_boring_prop (@{const "==>"} $ t $ u) = blanchet@49339: is_boring_prop t andalso is_boring_prop u blanchet@49339: | is_boring_prop (Const (@{const_name all}, _) $ (Abs (_, _, t)) $ u) = blanchet@49339: is_boring_prop t andalso is_boring_prop u blanchet@49339: | is_boring_prop (Const (@{const_name "=="}, _) $ t $ u) = blanchet@49339: is_boring_bool t andalso is_boring_bool u blanchet@49339: | is_boring_prop _ = true blanchet@49339: in blanchet@49339: is_boring_prop (prop_of th) andalso not (Thm.eq_thm_prop (@{thm ext}, th)) blanchet@49339: end blanchet@49339: blanchet@49339: fun theory_ord p = blanchet@49339: if Theory.eq_thy p then blanchet@49339: EQUAL blanchet@49339: else if Theory.subthy p then blanchet@49339: LESS blanchet@49339: else if Theory.subthy (swap p) then blanchet@49339: GREATER blanchet@49339: else case int_ord (pairself (length o Theory.ancestors_of) p) of blanchet@49339: EQUAL => string_ord (pairself Context.theory_name p) blanchet@49339: | order => order blanchet@49339: blanchet@49339: val thm_ord = theory_ord o pairself theory_of_thm blanchet@49339: blanchet@49341: val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}] blanchet@49341: blanchet@49333: fun interesting_terms_types_and_classes ctxt prover term_max_depth blanchet@49333: type_max_depth ts = blanchet@49266: let blanchet@49333: fun is_bad_const (x as (s, _)) args = blanchet@49333: member (op =) atp_logical_consts s orelse blanchet@49333: fst (is_built_in_const_for_prover ctxt prover x args) blanchet@49319: fun add_classes @{sort type} = I blanchet@49319: | add_classes S = union (op =) (map class_name_of S) blanchet@49266: fun do_add_type (Type (s, Ts)) = blanchet@49266: (not (member (op =) bad_types s) ? insert (op =) (type_name_of s)) blanchet@49266: #> fold do_add_type Ts blanchet@49319: | do_add_type (TFree (_, S)) = add_classes S blanchet@49319: | do_add_type (TVar (_, S)) = add_classes S blanchet@49266: fun add_type T = type_max_depth >= 0 ? do_add_type T blanchet@49266: fun mk_app s args = blanchet@49266: if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")" blanchet@49266: else s blanchet@49266: fun patternify ~1 _ = "" blanchet@49266: | patternify depth t = blanchet@49266: case strip_comb t of blanchet@49266: (Const (s, _), args) => blanchet@49266: mk_app (const_name_of s) (map (patternify (depth - 1)) args) blanchet@49266: | _ => "" blanchet@49266: fun add_term_patterns ~1 _ = I blanchet@49266: | add_term_patterns depth t = blanchet@49266: insert (op =) (patternify depth t) blanchet@49266: #> add_term_patterns (depth - 1) t blanchet@49266: val add_term = add_term_patterns term_max_depth blanchet@49266: fun add_patterns t = blanchet@49266: let val (head, args) = strip_comb t in blanchet@49266: (case head of blanchet@49333: Const (x as (_, T)) => blanchet@49333: not (is_bad_const x args) ? (add_term t #> add_type T) blanchet@49266: | Free (_, T) => add_type T blanchet@49266: | Var (_, T) => add_type T blanchet@49266: | Abs (_, T, body) => add_type T #> add_patterns body blanchet@49266: | _ => I) blanchet@49266: #> fold add_patterns args blanchet@49266: end blanchet@49341: in [] |> fold add_patterns ts end blanchet@49266: blanchet@49266: fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1}) blanchet@49266: blanchet@49312: val term_max_depth = 1 blanchet@49312: val type_max_depth = 1 blanchet@49266: blanchet@49266: (* TODO: Generate type classes for types? *) blanchet@49333: fun features_of ctxt prover thy status ts = blanchet@49318: thy_feature_name_of (Context.theory_name thy) :: blanchet@49333: interesting_terms_types_and_classes ctxt prover term_max_depth type_max_depth blanchet@49333: ts blanchet@49347: |> forall is_lambda_free ts ? cons "no_lams" blanchet@49347: |> forall (not o exists_Const is_exists) ts ? cons "no_skos" blanchet@49317: |> (case status of blanchet@49317: General => I blanchet@49317: | Induction => cons "induction" blanchet@49317: | Intro => cons "intro" blanchet@49317: | Inductive => cons "inductive" blanchet@49317: | Elim => cons "elim" blanchet@49317: | Simp => cons "simp" blanchet@49317: | Def => cons "def") blanchet@49266: blanchet@49341: fun isabelle_dependencies_of all_facts = thms_in_proof (SOME all_facts) blanchet@49266: blanchet@49266: val freezeT = Type.legacy_freeze_type blanchet@49266: blanchet@49266: fun freeze (t $ u) = freeze t $ freeze u blanchet@49266: | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t) blanchet@49266: | freeze (Var ((s, _), T)) = Free (s, freezeT T) blanchet@49266: | freeze (Const (s, T)) = Const (s, freezeT T) blanchet@49266: | freeze (Free (s, T)) = Free (s, freezeT T) blanchet@49266: | freeze t = t blanchet@49266: blanchet@49266: fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init blanchet@49266: blanchet@49336: fun run_prover_for_mash ctxt params prover facts goal = blanchet@49266: let blanchet@49266: val problem = blanchet@49266: {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1, blanchet@49304: facts = facts |> map (apfst (apfst (fn name => name ()))) blanchet@49333: |> map Untranslated_Fact} blanchet@49336: val prover = get_minimizing_prover ctxt Normal (K ()) prover blanchet@49266: in prover params (K (K (K ""))) problem end blanchet@49266: blanchet@49266: blanchet@49317: (*** Low-level communication with MaSh ***) blanchet@49317: blanchet@49338: fun write_file (xs, f) file = blanchet@49333: let val path = Path.explode file in blanchet@49338: File.write path ""; blanchet@49338: xs |> chunk_list 500 blanchet@49338: |> List.app (File.append path o space_implode "" o map f) blanchet@49333: end blanchet@49331: blanchet@49331: fun mash_info overlord = blanchet@49331: if overlord then (getenv "ISABELLE_HOME_USER", "") blanchet@49331: else (getenv "ISABELLE_TMP", serial_string ()) blanchet@49331: blanchet@49347: fun run_mash ctxt overlord (temp_dir, serial) core = blanchet@49326: let blanchet@49331: val log_file = temp_dir ^ "/mash_log" ^ serial blanchet@49331: val err_file = temp_dir ^ "/mash_err" ^ serial blanchet@49326: val command = blanchet@49331: mash_home () ^ "/mash.py --quiet --outputDir " ^ mash_state_dir () ^ blanchet@49331: " --log " ^ log_file ^ " " ^ core ^ " 2>&1 > " ^ err_file blanchet@49331: in blanchet@49331: trace_msg ctxt (fn () => "Running " ^ command); blanchet@49338: write_file ([], K "") log_file; blanchet@49338: write_file ([], K "") err_file; blanchet@49347: Isabelle_System.bash command; blanchet@49347: if overlord then () blanchet@49347: else (map (File.rm o Path.explode) [log_file, err_file]; ()); blanchet@49347: trace_msg ctxt (K "Done") blanchet@49331: end blanchet@49326: blanchet@49347: (* TODO: Eliminate code once "mash.py" handles sequences of ADD commands as fast blanchet@49347: as a single INIT. *) blanchet@49331: fun run_mash_init ctxt overlord write_access write_feats write_deps = blanchet@49331: let blanchet@49331: val info as (temp_dir, serial) = mash_info overlord blanchet@49331: val in_dir = temp_dir ^ "/mash_init" ^ serial blanchet@49347: val in_dir_path = in_dir |> Path.explode |> tap Isabelle_System.mkdir blanchet@49331: in blanchet@49331: write_file write_access (in_dir ^ "/mash_accessibility"); blanchet@49331: write_file write_feats (in_dir ^ "/mash_features"); blanchet@49331: write_file write_deps (in_dir ^ "/mash_dependencies"); blanchet@49347: run_mash ctxt overlord info ("--init --inputDir " ^ in_dir); blanchet@49347: (* FIXME: temporary hack *) blanchet@49347: if overlord then () blanchet@49347: else (Isabelle_System.bash ("rm -r -f " ^ File.shell_path in_dir_path); ()) blanchet@49331: end blanchet@49331: blanchet@49333: fun run_mash_commands ctxt overlord save max_suggs write_cmds read_suggs = blanchet@49331: let blanchet@49331: val info as (temp_dir, serial) = mash_info overlord blanchet@49333: val sugg_file = temp_dir ^ "/mash_suggs" ^ serial blanchet@49331: val cmd_file = temp_dir ^ "/mash_commands" ^ serial blanchet@49331: in blanchet@49338: write_file ([], K "") sugg_file; blanchet@49331: write_file write_cmds cmd_file; blanchet@49347: run_mash ctxt overlord info blanchet@49333: ("--inputFile " ^ cmd_file ^ " --predictions " ^ sugg_file ^ blanchet@49333: " --numberOfPredictions " ^ string_of_int max_suggs ^ blanchet@49331: (if save then " --saveModel" else "")); blanchet@49333: read_suggs (fn () => File.read_lines (Path.explode sugg_file)) blanchet@49347: |> tap (fn _ => blanchet@49347: if overlord then () blanchet@49347: else (map (File.rm o Path.explode) [sugg_file, cmd_file]; ())) blanchet@49331: end blanchet@49331: blanchet@49331: fun str_of_update (name, parents, feats, deps) = blanchet@49331: "! " ^ escape_meta name ^ ": " ^ escape_metas parents ^ "; " ^ blanchet@49326: escape_metas feats ^ "; " ^ escape_metas deps ^ "\n" blanchet@49326: blanchet@49331: fun str_of_query (parents, feats) = blanchet@49331: "? " ^ escape_metas parents ^ "; " ^ escape_metas feats blanchet@49331: blanchet@49331: fun init_str_of_update get (upd as (name, _, _, _)) = blanchet@49331: escape_meta name ^ ": " ^ escape_metas (get upd) ^ "\n" blanchet@49326: blanchet@49347: fun mash_CLEAR ctxt = blanchet@49329: let val path = mash_state_dir () |> Path.explode in blanchet@49347: trace_msg ctxt (K "MaSh CLEAR"); blanchet@49324: File.fold_dir (fn file => fn () => blanchet@49324: File.rm (Path.append path (Path.basic file))) blanchet@49324: path () blanchet@49324: end blanchet@49317: blanchet@49347: fun mash_INIT ctxt _ [] = mash_CLEAR ctxt blanchet@49331: | mash_INIT ctxt overlord upds = blanchet@49331: (trace_msg ctxt (fn () => "MaSh INIT " ^ blanchet@49331: elide_string 1000 (space_implode " " (map #1 upds))); blanchet@49338: run_mash_init ctxt overlord (upds, init_str_of_update #2) blanchet@49338: (upds, init_str_of_update #3) (upds, init_str_of_update #4)) blanchet@49317: blanchet@49331: fun mash_ADD _ _ [] = () blanchet@49331: | mash_ADD ctxt overlord upds = blanchet@49331: (trace_msg ctxt (fn () => "MaSh ADD " ^ blanchet@49331: elide_string 1000 (space_implode " " (map #1 upds))); blanchet@49338: run_mash_commands ctxt overlord true 0 (upds, str_of_update) (K ())) blanchet@49317: blanchet@49333: fun mash_QUERY ctxt overlord max_suggs (query as (_, feats)) = blanchet@49329: (trace_msg ctxt (fn () => "MaSh QUERY " ^ space_implode " " feats); blanchet@49333: run_mash_commands ctxt overlord false max_suggs blanchet@49338: ([query], str_of_query) blanchet@49333: (fn suggs => snd (extract_query (List.last (suggs ())))) blanchet@49326: handle List.Empty => []) blanchet@49317: blanchet@49317: blanchet@49266: (*** High-level communication with MaSh ***) blanchet@49266: blanchet@49336: fun try_graph ctxt when def f = blanchet@49336: f () blanchet@49336: handle Graph.CYCLES (cycle :: _) => blanchet@49336: (trace_msg ctxt (fn () => blanchet@49336: "Cycle involving " ^ commas cycle ^ " when " ^ when); def) blanchet@49336: | Graph.UNDEF name => blanchet@49336: (trace_msg ctxt (fn () => blanchet@49336: "Unknown fact " ^ quote name ^ " when " ^ when); def) blanchet@49336: blanchet@49316: type mash_state = blanchet@49331: {thys : bool Symtab.table, blanchet@49331: fact_graph : unit Graph.T} blanchet@49264: blanchet@49331: val empty_state = {thys = Symtab.empty, fact_graph = Graph.empty} blanchet@49316: blanchet@49316: local blanchet@49316: blanchet@49336: fun mash_load _ (state as (true, _)) = state blanchet@49336: | mash_load ctxt _ = blanchet@49324: let val path = mash_state_path () in blanchet@49317: (true, blanchet@49317: case try File.read_lines path of blanchet@49331: SOME (comp_thys :: incomp_thys :: fact_lines) => blanchet@49317: let blanchet@49331: fun add_thy comp thy = Symtab.update (thy, comp) blanchet@49337: fun add_edge_to name parent = blanchet@49337: Graph.default_node (parent, ()) blanchet@49337: #> Graph.add_edge (parent, name) blanchet@49331: fun add_fact_line line = blanchet@49331: case extract_query line of blanchet@49331: ("", _) => I (* shouldn't happen *) blanchet@49331: | (name, parents) => blanchet@49331: Graph.default_node (name, ()) blanchet@49337: #> fold (add_edge_to name) parents blanchet@49331: val thys = blanchet@49331: Symtab.empty |> fold (add_thy true) (unescape_metas comp_thys) blanchet@49331: |> fold (add_thy false) (unescape_metas incomp_thys) blanchet@49336: val fact_graph = blanchet@49337: try_graph ctxt "loading state" Graph.empty (fn () => blanchet@49336: Graph.empty |> fold add_fact_line fact_lines) blanchet@49331: in {thys = thys, fact_graph = fact_graph} end blanchet@49319: | _ => empty_state) blanchet@49317: end blanchet@49316: blanchet@49331: fun mash_save ({thys, fact_graph, ...} : mash_state) = blanchet@49316: let blanchet@49324: val path = mash_state_path () blanchet@49331: val thys = Symtab.dest thys blanchet@49333: val line_for_thys = escape_metas o AList.find (op =) thys blanchet@49333: fun fact_line_for name parents = blanchet@49333: escape_meta name ^ ": " ^ escape_metas parents blanchet@49331: val append_fact = File.append path o suffix "\n" oo fact_line_for blanchet@49316: in blanchet@49331: File.write path (line_for_thys true ^ "\n" ^ line_for_thys false ^ "\n"); blanchet@49331: Graph.fold (fn (name, ((), (parents, _))) => fn () => blanchet@49331: append_fact name (Graph.Keys.dest parents)) blanchet@49331: fact_graph () blanchet@49316: end blanchet@49316: blanchet@49317: val global_state = blanchet@49396: Synchronized.var "Sledgehammer_MaSh.global_state" (false, empty_state) blanchet@49316: blanchet@49316: in blanchet@49316: blanchet@49336: fun mash_map ctxt f = blanchet@49336: Synchronized.change global_state (mash_load ctxt ##> (f #> tap mash_save)) blanchet@49316: blanchet@49336: fun mash_get ctxt = blanchet@49336: Synchronized.change_result global_state (mash_load ctxt #> `snd) blanchet@49317: blanchet@49347: fun mash_unlearn ctxt = blanchet@49317: Synchronized.change global_state (fn _ => blanchet@49347: (mash_CLEAR ctxt; File.write (mash_state_path ()) ""; blanchet@49323: (true, empty_state))) blanchet@49316: blanchet@49316: end blanchet@49316: blanchet@49333: fun mash_could_suggest_facts () = mash_home () <> "" blanchet@49336: fun mash_can_suggest_facts ctxt = blanchet@49336: not (Graph.is_empty (#fact_graph (mash_get ctxt))) blanchet@49264: blanchet@49347: fun parents_wrt_facts facts fact_graph = blanchet@49331: let blanchet@49393: val facts = [] |> fold (cons o nickname_of o snd) facts blanchet@49345: val tab = Symtab.empty |> fold (fn name => Symtab.update (name, ())) facts blanchet@49392: fun insert_not_seen seen name = blanchet@49392: not (member (op =) seen name) ? insert (op =) name blanchet@49393: fun parents_of _ parents [] = parents blanchet@49392: | parents_of seen parents (name :: names) = blanchet@49347: if Symtab.defined tab name then blanchet@49392: parents_of (name :: seen) (name :: parents) names blanchet@49347: else blanchet@49392: parents_of (name :: seen) parents blanchet@49392: (Graph.Keys.fold (insert_not_seen seen) blanchet@49392: (Graph.imm_preds fact_graph name) names) blanchet@49392: in parents_of [] [] (Graph.maximals fact_graph) end blanchet@49331: blanchet@49333: (* Generate more suggestions than requested, because some might be thrown out blanchet@49333: later for various reasons and "meshing" gives better results with some blanchet@49333: slack. *) blanchet@49333: fun max_suggs_of max_facts = max_facts + Int.min (200, max_facts) blanchet@49333: blanchet@49335: fun is_fact_in_graph fact_graph (_, th) = blanchet@49393: can (Graph.get_node fact_graph) (nickname_of th) blanchet@49335: blanchet@49333: fun mash_suggest_facts ctxt ({overlord, ...} : params) prover max_facts hyp_ts blanchet@49333: concl_t facts = blanchet@49316: let blanchet@49317: val thy = Proof_Context.theory_of ctxt blanchet@49336: val fact_graph = #fact_graph (mash_get ctxt) blanchet@49347: val parents = parents_wrt_facts facts fact_graph blanchet@49333: val feats = features_of ctxt prover thy General (concl_t :: hyp_ts) blanchet@49333: val suggs = blanchet@49335: if Graph.is_empty fact_graph then [] blanchet@49335: else mash_QUERY ctxt overlord (max_suggs_of max_facts) (parents, feats) blanchet@49335: val selected = facts |> suggested_facts suggs blanchet@49335: val unknown = facts |> filter_out (is_fact_in_graph fact_graph) blanchet@49335: in (selected, unknown) end blanchet@49264: blanchet@49331: fun add_thys_for thy = blanchet@49331: let fun add comp thy = Symtab.update (Context.theory_name thy, comp) in blanchet@49331: add false thy #> fold (add true) (Theory.ancestors_of thy) blanchet@49331: end blanchet@49319: blanchet@49331: fun update_fact_graph ctxt (name, parents, feats, deps) (upds, graph) = blanchet@49331: let blanchet@49331: fun maybe_add_from from (accum as (parents, graph)) = blanchet@49336: try_graph ctxt "updating graph" accum (fn () => blanchet@49336: (from :: parents, Graph.add_edge_acyclic (from, name) graph)) blanchet@49336: val graph = graph |> Graph.default_node (name, ()) blanchet@49331: val (parents, graph) = ([], graph) |> fold maybe_add_from parents blanchet@49331: val (deps, graph) = ([], graph) |> fold maybe_add_from deps blanchet@49331: in ((name, parents, feats, deps) :: upds, graph) end blanchet@49321: blanchet@49333: val pass1_learn_timeout_factor = 0.5 blanchet@49333: blanchet@49347: (* Too many dependencies is a sign that a decision procedure is at work. There blanchet@49347: isn't much too learn from such proofs. *) blanchet@49347: val max_dependencies = 10 blanchet@49347: blanchet@49333: (* The timeout is understood in a very slack fashion. *) blanchet@49334: fun mash_learn_thy ctxt ({provers, verbose, overlord, ...} : params) thy timeout blanchet@49334: facts = blanchet@49319: let blanchet@49333: val timer = Timer.startRealTimer () blanchet@49333: val prover = hd provers blanchet@49333: fun timed_out frac = blanchet@49333: Time.> (Timer.checkRealTimer timer, time_mult frac timeout) blanchet@49336: val {fact_graph, ...} = mash_get ctxt blanchet@49335: val new_facts = blanchet@49335: facts |> filter_out (is_fact_in_graph fact_graph) blanchet@49335: |> sort (thm_ord o pairself snd) blanchet@49323: in blanchet@49323: if null new_facts then blanchet@49334: "" blanchet@49323: else blanchet@49319: let blanchet@49323: val ths = facts |> map snd blanchet@49330: val all_names = blanchet@49339: ths |> filter_out is_likely_tautology_or_too_meta blanchet@49393: |> map (rpair () o nickname_of) blanchet@49331: |> Symtab.make blanchet@49347: fun trim_deps deps = if length deps > max_dependencies then [] else deps blanchet@49333: fun do_fact _ (accum as (_, true)) = accum blanchet@49333: | do_fact ((_, (_, status)), th) ((parents, upds), false) = blanchet@49333: let blanchet@49393: val name = nickname_of th blanchet@49347: val feats = blanchet@49347: features_of ctxt prover (theory_of_thm th) status [prop_of th] blanchet@49347: val deps = isabelle_dependencies_of all_names th |> trim_deps blanchet@49333: val upd = (name, parents, feats, deps) blanchet@49333: in (([name], upd :: upds), timed_out pass1_learn_timeout_factor) end blanchet@49347: val parents = parents_wrt_facts facts fact_graph blanchet@49333: val ((_, upds), _) = blanchet@49333: ((parents, []), false) |> fold do_fact new_facts |>> apsnd rev blanchet@49333: val n = length upds blanchet@49331: fun trans {thys, fact_graph} = blanchet@49331: let blanchet@49331: val mash_INIT_or_ADD = blanchet@49331: if Graph.is_empty fact_graph then mash_INIT else mash_ADD blanchet@49331: val (upds, fact_graph) = blanchet@49331: ([], fact_graph) |> fold (update_fact_graph ctxt) upds blanchet@49331: in blanchet@49331: (mash_INIT_or_ADD ctxt overlord (rev upds); blanchet@49331: {thys = thys |> add_thys_for thy, blanchet@49331: fact_graph = fact_graph}) blanchet@49331: end blanchet@49333: in blanchet@49336: mash_map ctxt trans; blanchet@49334: if verbose then blanchet@49334: "Processed " ^ string_of_int n ^ " proof" ^ plural_s n ^ blanchet@49334: (if verbose then blanchet@49334: " in " ^ string_from_time (Timer.checkRealTimer timer) blanchet@49334: else blanchet@49334: "") ^ "." blanchet@49334: else blanchet@49334: "" blanchet@49333: end blanchet@49323: end blanchet@49319: blanchet@49336: fun mash_learn_proof ctxt ({provers, overlord, ...} : params) t facts used_ths = blanchet@49331: let blanchet@49331: val thy = Proof_Context.theory_of ctxt blanchet@49333: val prover = hd provers blanchet@49342: val name = ATP_Util.timestamp () ^ " " ^ serial_string () (* fresh enough *) blanchet@49333: val feats = features_of ctxt prover thy General [t] blanchet@49393: val deps = used_ths |> map nickname_of blanchet@49331: in blanchet@49336: mash_map ctxt (fn {thys, fact_graph} => blanchet@49324: let blanchet@49347: val parents = parents_wrt_facts facts fact_graph blanchet@49331: val upds = [(name, parents, feats, deps)] blanchet@49331: val (upds, fact_graph) = blanchet@49331: ([], fact_graph) |> fold (update_fact_graph ctxt) upds blanchet@49324: in blanchet@49331: mash_ADD ctxt overlord upds; blanchet@49331: {thys = thys, fact_graph = fact_graph} blanchet@49331: end) blanchet@49331: end blanchet@49264: blanchet@49333: (* The threshold should be large enough so that MaSh doesn't kick in for Auto blanchet@49333: Sledgehammer and Try. *) blanchet@49333: val min_secs_for_learning = 15 blanchet@49342: val learn_timeout_factor = 2.0 blanchet@49333: blanchet@49336: fun relevant_facts ctxt (params as {learn, fact_filter, timeout, ...}) prover blanchet@49336: max_facts ({add, only, ...} : fact_override) hyp_ts concl_t facts = blanchet@49329: if not (subset (op =) (the_list fact_filter, fact_filters)) then blanchet@49329: error ("Unknown fact filter: " ^ quote (the fact_filter) ^ ".") blanchet@49329: else if only then blanchet@49304: facts blanchet@49336: else if max_facts <= 0 orelse null facts then blanchet@49303: [] blanchet@49303: else blanchet@49303: let blanchet@49333: val thy = Proof_Context.theory_of ctxt blanchet@49342: fun maybe_learn () = blanchet@49336: if not learn orelse Async_Manager.has_running_threads MaShN then blanchet@49334: () blanchet@49334: else if Time.toSeconds timeout >= min_secs_for_learning then blanchet@49334: let blanchet@49342: val soft_timeout = time_mult learn_timeout_factor timeout blanchet@49342: val hard_timeout = time_mult 4.0 soft_timeout blanchet@49334: val birth_time = Time.now () blanchet@49334: val death_time = Time.+ (birth_time, hard_timeout) blanchet@49334: val desc = ("machine learner for Sledgehammer", "") blanchet@49334: in blanchet@49334: Async_Manager.launch MaShN birth_time death_time desc blanchet@49334: (fn () => blanchet@49334: (true, mash_learn_thy ctxt params thy soft_timeout facts)) blanchet@49334: end blanchet@49333: else blanchet@49333: () blanchet@49329: val fact_filter = blanchet@49329: case fact_filter of blanchet@49394: SOME ff => (() |> ff <> mepoN ? maybe_learn; ff) blanchet@49333: | NONE => blanchet@49342: if mash_can_suggest_facts ctxt then (maybe_learn (); meshN) blanchet@49394: else if mash_could_suggest_facts () then (maybe_learn (); mepoN) blanchet@49394: else mepoN blanchet@49303: val add_ths = Attrib.eval_thms ctxt add blanchet@49307: fun prepend_facts ths accepts = blanchet@49303: ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @ blanchet@49307: (accepts |> filter_out (member Thm.eq_thm_prop ths o snd))) blanchet@49308: |> take max_facts blanchet@49329: fun iter () = blanchet@49313: iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts blanchet@49313: concl_t facts blanchet@49329: fun mash () = blanchet@49335: mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts blanchet@49329: val mess = blanchet@49335: [] |> (if fact_filter <> mashN then cons (iter (), []) else I) blanchet@49394: |> (if fact_filter <> mepoN then cons (mash ()) else I) blanchet@49303: in blanchet@49328: mesh_facts max_facts mess blanchet@49303: |> not (null add_ths) ? prepend_facts add_ths blanchet@49303: end blanchet@49303: blanchet@49334: fun kill_learners () = Async_Manager.kill_threads MaShN "learner" blanchet@49334: fun running_learners () = Async_Manager.running_threads MaShN "learner" blanchet@49334: blanchet@49263: end;