src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML
author wenzelm
Wed, 29 Dec 2010 12:16:49 +0100
changeset 41655 2878845bc549
parent 41407 1e12d6495423
child 41711 d3be2a414d15
permissions -rw-r--r--
made SML/NJ happy;
more accurate dependencies;
     1 (*  Title:      HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML
     2     Author:     Jasmin Blanchette, TU Munich
     3 *)
     4 
     5 structure Mirabelle_Sledgehammer_Filter : MIRABELLE_ACTION =
     6 struct
     7 
     8 fun get args name default_value =
     9   case AList.lookup (op =) args name of
    10     SOME value => the (Real.fromString value)
    11   | NONE => default_value
    12 
    13 fun extract_relevance_fudge args
    14       {allow_ext, worse_irrel_freq, higher_order_irrel_weight, abs_rel_weight,
    15        abs_irrel_weight, skolem_irrel_weight, theory_const_rel_weight,
    16        theory_const_irrel_weight, intro_bonus, elim_bonus, simp_bonus,
    17        local_bonus, assum_bonus, chained_bonus, max_imperfect,
    18        max_imperfect_exp, threshold_divisor, ridiculous_threshold} =
    19   {allow_ext = allow_ext,
    20    worse_irrel_freq = get args "worse_irrel_freq" worse_irrel_freq,
    21    higher_order_irrel_weight =
    22        get args "higher_order_irrel_weight" higher_order_irrel_weight,
    23    abs_rel_weight = get args "abs_rel_weight" abs_rel_weight,
    24    abs_irrel_weight = get args "abs_irrel_weight" abs_irrel_weight,
    25    skolem_irrel_weight = get args "skolem_irrel_weight" skolem_irrel_weight,
    26    theory_const_rel_weight =
    27        get args "theory_const_rel_weight" theory_const_rel_weight,
    28    theory_const_irrel_weight =
    29        get args "theory_const_irrel_weight" theory_const_irrel_weight,
    30    intro_bonus = get args "intro_bonus" intro_bonus,
    31    elim_bonus = get args "elim_bonus" elim_bonus,
    32    simp_bonus = get args "simp_bonus" simp_bonus,
    33    local_bonus = get args "local_bonus" local_bonus,
    34    assum_bonus = get args "assum_bonus" assum_bonus,
    35    chained_bonus = get args "chained_bonus" chained_bonus,
    36    max_imperfect = get args "max_imperfect" max_imperfect,
    37    max_imperfect_exp = get args "max_imperfect_exp" max_imperfect_exp,
    38    threshold_divisor = get args "threshold_divisor" threshold_divisor,
    39    ridiculous_threshold = get args "ridiculous_threshold" ridiculous_threshold}
    40 
    41 structure Prooftab =
    42   Table(type key = int * int val ord = prod_ord int_ord int_ord)
    43 
    44 val proof_table = Unsynchronized.ref (Prooftab.empty: string list list Prooftab.table)
    45 
    46 val num_successes = Unsynchronized.ref ([] : (int * int) list)
    47 val num_failures = Unsynchronized.ref ([] : (int * int) list)
    48 val num_found_proofs = Unsynchronized.ref ([] : (int * int) list)
    49 val num_lost_proofs = Unsynchronized.ref ([] : (int * int) list)
    50 val num_found_facts = Unsynchronized.ref ([] : (int * int) list)
    51 val num_lost_facts = Unsynchronized.ref ([] : (int * int) list)
    52 
    53 fun get id c = the_default 0 (AList.lookup (op =) (!c) id)
    54 fun add id c n =
    55   c := (case AList.lookup (op =) (!c) id of
    56           SOME m => AList.update (op =) (id, m + n) (!c)
    57         | NONE => (id, n) :: !c)
    58 
    59 fun init proof_file _ thy =
    60   let
    61     fun do_line line =
    62       case line |> space_explode ":" of
    63         [line_num, col_num, proof] =>
    64         SOME (pairself (the o Int.fromString) (line_num, col_num),
    65               proof |> space_explode " " |> filter_out (curry (op =) ""))
    66        | _ => NONE
    67     val proofs = File.read (Path.explode proof_file)
    68     val proof_tab =
    69       proofs |> space_explode "\n"
    70              |> map_filter do_line
    71              |> AList.coalesce (op =)
    72              |> Prooftab.make
    73   in proof_table := proof_tab; thy end
    74 
    75 fun percentage a b = if b = 0 then "N/A" else string_of_int (a * 100 div b)
    76 fun percentage_alt a b = percentage a (a + b)
    77 
    78 fun done id ({log, ...} : Mirabelle.done_args) =
    79   if get id num_successes + get id num_failures > 0 then
    80     (log "";
    81      log ("Number of overall successes: " ^
    82           string_of_int (get id num_successes));
    83      log ("Number of overall failures: " ^ string_of_int (get id num_failures));
    84      log ("Overall success rate: " ^
    85           percentage_alt (get id num_successes) (get id num_failures) ^ "%");
    86      log ("Number of found proofs: " ^ string_of_int (get id num_found_proofs));
    87      log ("Number of lost proofs: " ^ string_of_int (get id num_lost_proofs));
    88      log ("Proof found rate: " ^
    89           percentage_alt (get id num_found_proofs) (get id num_lost_proofs) ^
    90           "%");
    91      log ("Number of found facts: " ^ string_of_int (get id num_found_facts));
    92      log ("Number of lost facts: " ^ string_of_int (get id num_lost_facts));
    93      log ("Fact found rate: " ^
    94           percentage_alt (get id num_found_facts) (get id num_lost_facts) ^
    95           "%"))
    96   else
    97     ()
    98 
    99 val default_prover = ATP_Systems.eN (* arbitrary ATP *)
   100 
   101 fun with_index (i, s) = s ^ "@" ^ string_of_int i
   102 
   103 fun action args id ({pre, pos, log, ...} : Mirabelle.run_args) =
   104   case (Position.line_of pos, Position.column_of pos) of
   105     (SOME line_num, SOME col_num) =>
   106     (case Prooftab.lookup (!proof_table) (line_num, col_num) of
   107        SOME proofs =>
   108        let
   109          val {context = ctxt, facts, goal} = Proof.goal pre
   110          val prover = AList.lookup (op =) args "prover"
   111                       |> the_default default_prover
   112          val default_max_relevant =
   113            Sledgehammer_Provers.default_max_relevant_for_prover ctxt prover
   114         val is_built_in_const =
   115           Sledgehammer_Provers.is_built_in_const_for_prover ctxt default_prover
   116          val relevance_fudge =
   117            extract_relevance_fudge args
   118                (Sledgehammer_Provers.relevance_fudge_for_prover ctxt prover)
   119          val relevance_override = {add = [], del = [], only = false}
   120          val {relevance_thresholds, type_sys, max_relevant, ...} =
   121            Sledgehammer_Isar.default_params ctxt args
   122          val subgoal = 1
   123          val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal subgoal
   124          val no_dangerous_types =
   125            Sledgehammer_ATP_Translate.types_dangerous_types type_sys
   126          val facts =
   127            Sledgehammer_Filter.relevant_facts ctxt no_dangerous_types
   128                relevance_thresholds
   129                (the_default default_max_relevant max_relevant) is_built_in_const
   130                relevance_fudge relevance_override facts hyp_ts concl_t
   131            |> map (fst o fst)
   132          val (found_facts, lost_facts) =
   133            List.concat proofs |> sort_distinct string_ord
   134            |> map (fn fact => (find_index (curry (op =) fact) facts, fact))
   135            |> List.partition (curry (op <=) 0 o fst)
   136            |>> sort (prod_ord int_ord string_ord) ||> map snd
   137          val found_proofs = filter (forall (member (op =) facts)) proofs
   138          val n = length found_proofs
   139          val _ =
   140            if n = 0 then
   141              (add id num_failures 1; log "Failure")
   142            else
   143              (add id num_successes 1;
   144               add id num_found_proofs n;
   145               log ("Success (" ^ string_of_int n ^ " of " ^
   146                    string_of_int (length proofs) ^ " proofs)"))
   147          val _ = add id num_lost_proofs (length proofs - n)
   148          val _ = add id num_found_facts (length found_facts)
   149          val _ = add id num_lost_facts (length lost_facts)
   150          val _ =
   151            if null found_facts then
   152              ()
   153            else
   154              let
   155                val found_weight =
   156                  Real.fromInt (fold (fn (n, _) =>
   157                                         Integer.add (n * n)) found_facts 0)
   158                    / Real.fromInt (length found_facts)
   159                  |> Math.sqrt |> Real.ceil
   160              in
   161                log ("Found facts (among " ^ string_of_int (length facts) ^
   162                     ", weight " ^ string_of_int found_weight ^ "): " ^
   163                     commas (map with_index found_facts))
   164              end
   165          val _ = if null lost_facts then
   166                    ()
   167                  else
   168                    log ("Lost facts (among " ^ string_of_int (length facts) ^
   169                         "): " ^ commas lost_facts)
   170        in () end
   171      | NONE => log "No known proof")
   172   | _ => ()
   173 
   174 val proof_fileK = "proof_file"
   175 
   176 fun invoke args =
   177   let
   178     val (pf_args, other_args) =
   179       args |> List.partition (curry (op =) proof_fileK o fst)
   180     val proof_file = case pf_args of
   181                        [] => error "No \"proof_file\" specified"
   182                      | (_, s) :: _ => s
   183   in Mirabelle.register (init proof_file, action other_args, done) end
   184 
   185 end;
   186 
   187 (* Workaround to keep the "mirabelle.pl" script happy *)
   188 structure Mirabelle_Sledgehammer_filter = Mirabelle_Sledgehammer_Filter;