src/HOL/TPTP/mash_eval.ML
author blanchet
Mon, 23 Jul 2012 15:32:30 +0200
changeset 49453 3e45c98fe127
parent 49421 b002cc16aa99
child 49545 d443166f9520
permissions -rw-r--r--
distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
     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_Fact
    20 open Sledgehammer_MaSh
    21 
    22 val MePoN = "MePo"
    23 val MaShN = "MaSh"
    24 val MeshN = "Mesh"
    25 val IsarN = "Isar"
    26 
    27 val max_facts_slack = 2
    28 
    29 val all_names = map (rpair () o nickname_of) #> Symtab.make
    30 
    31 fun evaluate_mash_suggestions ctxt params thy file_name =
    32   let
    33     val {provers, max_facts, slice, type_enc, lam_trans, timeout, ...} =
    34       Sledgehammer_Isar.default_params ctxt []
    35     val prover = hd provers
    36     val slack_max_facts = max_facts_slack * the max_facts
    37     val path = file_name |> Path.explode
    38     val lines = path |> File.read_lines
    39     val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
    40     val facts = all_facts_of (Proof_Context.init_global thy) css_table
    41     val all_names = all_names (facts |> map snd)
    42     val mepo_ok = Unsynchronized.ref 0
    43     val mash_ok = Unsynchronized.ref 0
    44     val mesh_ok = Unsynchronized.ref 0
    45     val isar_ok = Unsynchronized.ref 0
    46     fun with_index facts s = (find_index (curry (op =) s) facts + 1, s)
    47     fun index_string (j, s) = s ^ "@" ^ string_of_int j
    48     fun str_of_res label facts {outcome, run_time, used_facts, ...} =
    49       let val facts = facts |> map (fn ((name, _), _) => name ()) in
    50         "  " ^ label ^ ": " ^
    51         (if is_none outcome then
    52            "Success (" ^ ATP_Util.string_from_time run_time ^ "): " ^
    53            (used_facts |> map (with_index facts o fst)
    54                        |> sort (int_ord o pairself fst)
    55                        |> map index_string
    56                        |> space_implode " ") ^
    57            (if length facts < the max_facts then
    58               " (of " ^ string_of_int (length facts) ^ ")"
    59             else
    60               "")
    61          else
    62            "Failure: " ^
    63            (facts |> take (the max_facts) |> tag_list 1
    64                   |> map index_string
    65                   |> space_implode " "))
    66       end
    67     fun solve_goal (j, line) =
    68       let
    69         val (name, suggs) = extract_query line
    70         val th =
    71           case find_first (fn (_, th) => nickname_of th = name) facts of
    72             SOME (_, th) => th
    73           | NONE => error ("No fact called \"" ^ name ^ "\"")
    74         val goal = goal_of_thm thy th
    75         val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
    76         val isar_deps = isar_dependencies_of all_names th |> these
    77         val facts = facts |> filter (fn (_, th') => thm_ord (th', th) = LESS)
    78         val mepo_facts =
    79           Sledgehammer_MePo.mepo_suggested_facts ctxt params prover
    80               slack_max_facts NONE hyp_ts concl_t facts
    81           |> Sledgehammer_MePo.weight_mepo_facts
    82         val mash_facts = suggested_facts suggs facts
    83         val mess = [(mepo_facts, []), (mash_facts, [])]
    84         val mesh_facts = mesh_facts slack_max_facts mess
    85         val isar_facts = suggested_facts (map (rpair 1.0) isar_deps) facts
    86         fun prove ok heading get facts =
    87           let
    88             val facts =
    89               facts |> map get
    90                     |> Sledgehammer_Fact.maybe_instantiate_inducts ctxt hyp_ts
    91                                                                    concl_t
    92                     |> take (the max_facts)
    93             val res as {outcome, ...} =
    94               run_prover_for_mash ctxt params prover facts goal
    95             val _ = if is_none outcome then ok := !ok + 1 else ()
    96           in str_of_res heading facts res end
    97         val mepo_s = prove mepo_ok MePoN fst mepo_facts
    98         val mash_s = prove mash_ok MaShN fst mash_facts
    99         val mesh_s = prove mesh_ok MeshN I mesh_facts
   100         val isar_s = prove isar_ok IsarN fst isar_facts
   101       in
   102         ["Goal " ^ string_of_int j ^ ": " ^ name, mepo_s, mash_s, mesh_s,
   103          isar_s]
   104         |> cat_lines |> tracing
   105       end
   106     fun total_of heading ok n =
   107       "  " ^ heading ^ ": " ^ string_of_int (!ok) ^ " (" ^
   108       Real.fmt (StringCvt.FIX (SOME 1))
   109                (100.0 * Real.fromInt (!ok) / Real.fromInt n) ^ "%)"
   110     val inst_inducts = Config.get ctxt Sledgehammer_Fact.instantiate_inducts
   111     val options =
   112       [prover, string_of_int (the max_facts) ^ " facts",
   113        "slice" |> not slice ? prefix "dont_", the_default "smart" type_enc,
   114        the_default "smart" lam_trans, ATP_Util.string_from_time timeout,
   115        "instantiate_inducts" |> not inst_inducts ? prefix "dont_"]
   116     val n = length lines
   117   in
   118     tracing " * * *";
   119     tracing ("Options: " ^ commas options);
   120     List.app solve_goal (tag_list 1 lines);
   121     ["Successes (of " ^ string_of_int n ^ " goals)",
   122      total_of MePoN mepo_ok n,
   123      total_of MaShN mash_ok n,
   124      total_of MeshN mesh_ok n,
   125      total_of IsarN isar_ok n]
   126     |> cat_lines |> tracing
   127   end
   128 
   129 end;