src/HOL/TPTP/mash_eval.ML
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 49318 f1d135d0ea69
parent 49315 9910021c80a7
child 49322 7c78f14d20cf
permissions -rw-r--r--
improved MaSh string escaping and make more operations string-based
     1 (*  Title:      HOL/TPTP/mash_eval.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2012
     4 
     5 Evaluate proof suggestions from MaSh (Machine-learning for Sledgehammer).
     6 *)
     7 
     8 signature MASH_EVAL =
     9 sig
    10   type params = Sledgehammer_Provers.params
    11 
    12   val evaluate_mash_suggestions :
    13     Proof.context -> params -> theory -> string -> unit
    14 end;
    15 
    16 structure MaSh_Eval : MASH_EVAL =
    17 struct
    18 
    19 open Sledgehammer_Filter_MaSh
    20 
    21 val unescape_meta =
    22   let
    23     fun un accum [] = String.implode (rev accum)
    24       | un accum (#"\\" :: d1 :: d2 :: d3 :: cs) =
    25         (case Int.fromString (String.implode [d1, d2, d3]) of
    26            SOME n => un (Char.chr n :: accum) cs
    27          | NONE => "" (* error *))
    28       | un _ (#"\\" :: _) = "" (* error *)
    29       | un accum (c :: cs) = un (c :: accum) cs
    30   in un [] o String.explode end
    31 
    32 val isarN = "Isa"
    33 val iterN = "Iter"
    34 val mashN = "MaSh"
    35 val meshN = "Mesh"
    36 
    37 val max_facts_slack = 2
    38 
    39 fun evaluate_mash_suggestions ctxt params thy file_name =
    40   let
    41     val {provers, max_facts, slice, type_enc, lam_trans, timeout, ...} =
    42       Sledgehammer_Isar.default_params ctxt []
    43     val prover_name = hd provers
    44     val path = file_name |> Path.explode
    45     val lines = path |> File.read_lines
    46     val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
    47     val facts = all_non_tautological_facts_of thy css_table
    48     val all_names = facts |> map (Thm.get_name_hint o snd)
    49     val iter_ok = Unsynchronized.ref 0
    50     val mash_ok = Unsynchronized.ref 0
    51     val mesh_ok = Unsynchronized.ref 0
    52     val isar_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)
    57       #> take (max_facts_slack * the max_facts)
    58       #> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts concl_t
    59     fun mesh_facts fsp =
    60       let
    61         val (fs1, fs2) =
    62           fsp |> pairself (map (apfst (apfst (fn name => name ()))))
    63         val fact_eq = (op =) o pairself fst
    64         fun score_in f fs =
    65           case find_index (curry fact_eq f) fs of
    66             ~1 => length fs
    67           | j => j
    68         fun score_of f = score_in f fs1 + score_in f fs2
    69       in
    70         union fact_eq fs1 fs2
    71         |> map (`score_of) |> sort (int_ord o pairself fst) |> map snd
    72         |> take (the max_facts)
    73         |> map (apfst (apfst K))
    74       end
    75     fun with_index facts s = (find_index (curry (op =) s) facts + 1, s)
    76     fun index_string (j, s) = s ^ "@" ^ string_of_int j
    77     fun str_of_res label facts {outcome, run_time, used_facts, ...} =
    78       let val facts = facts |> map (fn ((name, _), _) => name ()) in
    79         "  " ^ label ^ ": " ^
    80         (if is_none outcome then
    81            "Success (" ^ ATP_Util.string_from_time run_time ^ "): " ^
    82            (used_facts |> map (with_index facts o fst)
    83                        |> sort (int_ord o pairself fst)
    84                        |> map index_string
    85                        |> space_implode " ") ^
    86            (if length facts < the max_facts then
    87               " (of " ^ string_of_int (length facts) ^ ")"
    88             else
    89               "")
    90          else
    91            "Failure: " ^
    92            (facts |> take (the max_facts) |> tag_list 1
    93                   |> map index_string
    94                   |> space_implode " "))
    95       end
    96     fun prove ok heading facts goal =
    97       let
    98         val facts = facts |> take (the max_facts)
    99         val res as {outcome, ...} = run_prover ctxt params facts goal
   100         val _ = if is_none outcome then ok := !ok + 1 else ()
   101       in str_of_res heading facts res end
   102     fun solve_goal j name suggs =
   103       let
   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 isar_deps = isabelle_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 isar_facts = sugg_facts hyp_ts concl_t facts isar_deps
   113         val iter_facts =
   114           Sledgehammer_Filter_Iter.iterative_relevant_facts ctxt params
   115               prover_name (max_facts_slack * the max_facts) NONE hyp_ts concl_t
   116               facts
   117         val mash_facts = sugg_facts hyp_ts concl_t facts suggs
   118         val mesh_facts = mesh_facts (iter_facts, mash_facts)
   119         val iter_s = prove iter_ok iterN iter_facts goal
   120         val mash_s = prove mash_ok mashN mash_facts goal
   121         val mesh_s = prove mesh_ok meshN mesh_facts goal
   122         val isar_s = prove isar_ok isarN isar_facts goal
   123       in
   124         ["Goal " ^ string_of_int j ^ ": " ^ name, iter_s, mash_s, mesh_s,
   125          isar_s]
   126         |> cat_lines |> tracing
   127       end
   128     val explode_suggs =
   129       space_explode " " #> filter_out (curry (op =) "") #> map unescape_meta
   130     fun do_line (j, line) =
   131       case space_explode ":" line of
   132         [goal_name, suggs] =>
   133         solve_goal j (unescape_meta goal_name) (explode_suggs suggs)
   134       | _ => ()
   135     fun total_of heading ok n =
   136       " " ^ heading ^ ": " ^ string_of_int (!ok) ^ " (" ^
   137       Real.fmt (StringCvt.FIX (SOME 1))
   138                (100.0 * Real.fromInt (!ok) / Real.fromInt n) ^ "%)"
   139     val inst_inducts = Config.get ctxt Sledgehammer_Fact.instantiate_inducts
   140     val options =
   141       [prover_name, string_of_int (the max_facts) ^ " facts",
   142        "slice" |> not slice ? prefix "dont_", the_default "smart" type_enc,
   143        the_default "smart" lam_trans, ATP_Util.string_from_time timeout,
   144        "instantiate_inducts" |> not inst_inducts ? prefix "dont_"]
   145     val n = length lines
   146   in
   147     tracing " * * *";
   148     tracing ("Options: " ^ commas options);
   149     List.app do_line (tag_list 1 lines);
   150     ["Successes (of " ^ string_of_int n ^ " goals)",
   151      total_of iterN iter_ok n,
   152      total_of mashN mash_ok n,
   153      total_of meshN mesh_ok n,
   154      total_of isarN isar_ok n]
   155     |> cat_lines |> tracing
   156   end
   157 
   158 end;