src/HOL/TPTP/mash_import.ML
author blanchet
Wed, 11 Jul 2012 21:43:19 +0200
changeset 49265 1065c307fafe
parent 49260 854a47677335
child 49266 6cdcfbddc077
permissions -rw-r--r--
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
     1 (*  Title:      HOL/TPTP/mash_import.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2012
     4 
     5 Import proof suggestions from MaSh (Machine-learning for Sledgehammer) and
     6 evaluate them.
     7 *)
     8 
     9 signature MASH_IMPORT =
    10 sig
    11   val import_and_evaluate_mash_suggestions :
    12     Proof.context -> theory -> string -> unit
    13 end;
    14 
    15 structure MaSh_Import : MASH_IMPORT =
    16 struct
    17 
    18 open MaSh_Export
    19 
    20 val unescape_meta =
    21   let
    22     fun un accum [] = String.implode (rev accum)
    23       | un accum (#"\\" :: d1 :: d2 :: d3 :: cs) =
    24         (case Int.fromString (String.implode [d1, d2, d3]) of
    25            SOME n => un (Char.chr n :: accum) cs
    26          | NONE => "" (* error *))
    27       | un _ (#"\\" :: _) = "" (* error *)
    28       | un accum (c :: cs) = un (c :: accum) cs
    29   in un [] o String.explode end
    30 
    31 val of_fact_name = unescape_meta
    32 
    33 val depN = "Dependencies"
    34 val mengN = "Meng & Paulson"
    35 val mashN = "MaSh"
    36 val meng_mashN = "M&P+MaSh"
    37 
    38 val max_relevant_slack = 2
    39 
    40 fun import_and_evaluate_mash_suggestions ctxt thy file_name =
    41   let
    42     val {provers, max_relevant, slice, type_enc, lam_trans, timeout, ...} =
    43       Sledgehammer_Isar.default_params ctxt []
    44     val prover_name = hd provers
    45     val path = file_name |> Path.explode
    46     val lines = path |> File.read_lines
    47     val facts = non_tautological_facts_of thy
    48     val all_names = facts |> map (Thm.get_name_hint o snd)
    49     val meng_ok = Unsynchronized.ref 0
    50     val mash_ok = Unsynchronized.ref 0
    51     val meng_mash_ok = Unsynchronized.ref 0
    52     val dep_ok = Unsynchronized.ref 0
    53     fun find_sugg facts sugg =
    54       find_first (fn (_, th) => Thm.get_name_hint th = sugg) facts
    55     fun sugg_facts hyp_ts concl_t facts =
    56       map_filter (find_sugg facts o of_fact_name)
    57       #> take (max_relevant_slack * the max_relevant)
    58       #> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts concl_t
    59       #> map (apfst (apfst (fn name => name ())))
    60     fun meng_mash_facts fs1 fs2 =
    61       let
    62         val fact_eq = (op =) o pairself fst
    63         fun score_in f fs =
    64           case find_index (curry fact_eq f) fs of
    65             ~1 => length fs
    66           | j => j
    67         fun score_of f = score_in f fs1 + score_in f fs2
    68       in
    69         union fact_eq fs1 fs2
    70         |> map (`score_of) |> sort (int_ord o pairself fst) |> map snd
    71         |> take (the max_relevant)
    72       end
    73     fun with_index facts s =
    74       (find_index (fn ((s', _), _) => s = s') facts + 1, s)
    75     fun index_string (j, s) = s ^ "@" ^ string_of_int j
    76     fun str_of_res label facts {outcome, run_time, used_facts, ...} =
    77       "  " ^ label ^ ": " ^
    78       (if is_none outcome then
    79          "Success (" ^ ATP_Util.string_from_time run_time ^ "): " ^
    80          (used_facts |> map (with_index facts o fst)
    81                      |> sort (int_ord o pairself fst)
    82                      |> map index_string
    83                      |> space_implode " ") ^
    84          (if length facts < the max_relevant then
    85             " (of " ^ string_of_int (length facts) ^ ")"
    86           else
    87             "")
    88        else
    89          "Failure: " ^
    90          (facts |> map (fst o fst)
    91                 |> take (the max_relevant)
    92                 |> tag_list 1
    93                 |> map index_string
    94                 |> space_implode " "))
    95     fun run_sh ok heading facts goal =
    96       let
    97         val facts = facts |> take (the max_relevant)
    98         val res as {outcome, ...} = run_sledgehammer ctxt facts goal
    99         val _ = if is_none outcome then ok := !ok + 1 else ()
   100       in str_of_res heading facts res end
   101     fun solve_goal j name suggs =
   102       let
   103         val name = of_fact_name name
   104         val th =
   105           case find_first (fn (_, th) => Thm.get_name_hint th = name) facts of
   106             SOME (_, th) => th
   107           | NONE => error ("No fact called \"" ^ name ^ "\"")
   108         val deps = dependencies_of all_names th
   109         val goal = goal_of_thm thy th
   110         val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
   111         val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
   112         val deps_facts = sugg_facts hyp_ts concl_t facts deps
   113         val meng_facts =
   114           meng_paulson_facts ctxt (max_relevant_slack * the max_relevant) goal
   115                              facts
   116         val mash_facts = sugg_facts hyp_ts concl_t facts suggs
   117         val meng_mash_facts = meng_mash_facts meng_facts mash_facts
   118         val meng_s = run_sh meng_ok mengN meng_facts goal
   119         val mash_s = run_sh mash_ok mashN mash_facts goal
   120         val meng_mash_s = run_sh meng_mash_ok meng_mashN meng_mash_facts goal
   121         val dep_s = run_sh dep_ok depN deps_facts goal
   122       in
   123         ["Goal " ^ string_of_int j ^ ": " ^ name, meng_s, mash_s, meng_mash_s,
   124          dep_s]
   125         |> cat_lines |> tracing
   126       end
   127     val explode_suggs = space_explode " " #> filter_out (curry (op =) "")
   128     fun do_line (j, line) =
   129       case space_explode ":" line of
   130         [goal_name, suggs] => solve_goal j goal_name (explode_suggs suggs)
   131       | _ => ()
   132     fun total_of heading ok n =
   133       " " ^ heading ^ ": " ^ string_of_int (!ok) ^ " (" ^
   134       Real.fmt (StringCvt.FIX (SOME 1))
   135                (100.0 * Real.fromInt (!ok) / Real.fromInt n) ^ "%)"
   136     val inst_inducts = Config.get ctxt Sledgehammer_Fact.instantiate_inducts
   137     val options =
   138       [prover_name, string_of_int (the max_relevant) ^ " facts",
   139        "slice" |> not slice ? prefix "dont_", the_default "smart" type_enc,
   140        the_default "smart" lam_trans, ATP_Util.string_from_time timeout,
   141        "instantiate_inducts" |> not inst_inducts ? prefix "dont_"]
   142     val n = length lines
   143   in
   144     tracing " * * *";
   145     tracing ("Options: " ^ commas options);
   146     List.app do_line (tag_list 1 lines);
   147     ["Successes (of " ^ string_of_int n ^ " goals)",
   148      total_of mengN meng_ok n,
   149      total_of mashN mash_ok n,
   150      total_of meng_mashN meng_mash_ok n,
   151      total_of depN dep_ok n]
   152     |> cat_lines |> tracing
   153   end
   154 
   155 end;