src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML
author nipkow
Sun, 28 Oct 2018 11:00:00 +0100
changeset 69200 f2bb47056d8f
parent 63806 c54a53ef1873
permissions -rw-r--r--
tuned names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47847
7cddb6c8f93c updated headers;
wenzelm
parents: 47730
diff changeset
     1
(*  Title:      HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Munich
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
     3
*)
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
     4
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
     5
structure Mirabelle_Sledgehammer_Filter : MIRABELLE_ACTION =
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
     6
struct
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
     7
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
     8
fun get args name default_value =
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
     9
  case AList.lookup (op =) args name of
63806
c54a53ef1873 clarified modules;
wenzelm
parents: 59058
diff changeset
    10
    SOME value => Value.parse_real value
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    11
  | NONE => default_value
38992
542474156c66 introduce fudge factors to deal with "theory const"
blanchet
parents: 38988
diff changeset
    12
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    13
fun extract_relevance_fudge args
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    14
      {local_const_multiplier, worse_irrel_freq, higher_order_irrel_weight, abs_rel_weight,
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    15
       abs_irrel_weight, theory_const_rel_weight, theory_const_irrel_weight,
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    16
       chained_const_irrel_weight, intro_bonus, elim_bonus, simp_bonus, local_bonus, assum_bonus,
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    17
       chained_bonus, max_imperfect, max_imperfect_exp, threshold_divisor, ridiculous_threshold} =
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    18
  {local_const_multiplier = get args "local_const_multiplier" local_const_multiplier,
41159
1e12d6495423 honor "overlord" option for SMT solvers as well and don't pass "ext" to them
blanchet
parents: 41138
diff changeset
    19
   worse_irrel_freq = get args "worse_irrel_freq" worse_irrel_freq,
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    20
   higher_order_irrel_weight = get args "higher_order_irrel_weight" higher_order_irrel_weight,
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    21
   abs_rel_weight = get args "abs_rel_weight" abs_rel_weight,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    22
   abs_irrel_weight = get args "abs_irrel_weight" abs_irrel_weight,
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    23
   theory_const_rel_weight = get args "theory_const_rel_weight" theory_const_rel_weight,
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    24
   theory_const_irrel_weight = get args "theory_const_irrel_weight" theory_const_irrel_weight,
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    25
   chained_const_irrel_weight = get args "chained_const_irrel_weight" chained_const_irrel_weight,
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    26
   intro_bonus = get args "intro_bonus" intro_bonus,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    27
   elim_bonus = get args "elim_bonus" elim_bonus,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    28
   simp_bonus = get args "simp_bonus" simp_bonus,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    29
   local_bonus = get args "local_bonus" local_bonus,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    30
   assum_bonus = get args "assum_bonus" assum_bonus,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    31
   chained_bonus = get args "chained_bonus" chained_bonus,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    32
   max_imperfect = get args "max_imperfect" max_imperfect,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    33
   max_imperfect_exp = get args "max_imperfect_exp" max_imperfect_exp,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    34
   threshold_divisor = get args "threshold_divisor" threshold_divisor,
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
    35
   ridiculous_threshold = get args "ridiculous_threshold" ridiculous_threshold}
38902
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
    36
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    37
structure Prooftab =
41407
2878845bc549 made SML/NJ happy;
wenzelm
parents: 41159
diff changeset
    38
  Table(type key = int * int val ord = prod_ord int_ord int_ord)
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    39
41407
2878845bc549 made SML/NJ happy;
wenzelm
parents: 41159
diff changeset
    40
val proof_table = Unsynchronized.ref (Prooftab.empty: string list list Prooftab.table)
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    41
38896
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    42
val num_successes = Unsynchronized.ref ([] : (int * int) list)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    43
val num_failures = Unsynchronized.ref ([] : (int * int) list)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    44
val num_found_proofs = Unsynchronized.ref ([] : (int * int) list)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    45
val num_lost_proofs = Unsynchronized.ref ([] : (int * int) list)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    46
val num_found_facts = Unsynchronized.ref ([] : (int * int) list)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    47
val num_lost_facts = Unsynchronized.ref ([] : (int * int) list)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    48
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    49
fun get id c = the_default 0 (AList.lookup (op =) (!c) id)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    50
fun add id c n =
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    51
  c := (case AList.lookup (op =) (!c) id of
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    52
         SOME m => AList.update (op =) (id, m + n) (!c)
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    53
       | NONE => (id, n) :: !c)
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    54
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    55
fun init proof_file _ thy =
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    56
  let
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    57
    fun do_line line =
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    58
      (case line |> space_explode ":" of
43710
7270ae921cf2 discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents: 43351
diff changeset
    59
        [line_num, offset, proof] =>
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58928
diff changeset
    60
        SOME (apply2 (the o Int.fromString) (line_num, offset),
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    61
              proof |> space_explode " " |> filter_out (curry (op =) ""))
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    62
       | _ => NONE)
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    63
    val proofs = File.read (Path.explode proof_file)
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    64
    val proof_tab =
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    65
      proofs |> space_explode "\n"
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    66
             |> map_filter do_line
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    67
             |> AList.coalesce (op =)
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    68
             |> Prooftab.make
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    69
  in proof_table := proof_tab; thy end
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    70
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    71
fun percentage a b = if b = 0 then "N/A" else string_of_int (a * 100 div b)
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    72
fun percentage_alt a b = percentage a (a + b)
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    73
38896
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    74
fun done id ({log, ...} : Mirabelle.done_args) =
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    75
  if get id num_successes + get id num_failures > 0 then
38897
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
    76
    (log "";
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    77
     log ("Number of overall successes: " ^ string_of_int (get id num_successes));
38897
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
    78
     log ("Number of overall failures: " ^ string_of_int (get id num_failures));
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    79
     log ("Overall success rate: " ^
38896
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    80
          percentage_alt (get id num_successes) (get id num_failures) ^ "%");
38897
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
    81
     log ("Number of found proofs: " ^ string_of_int (get id num_found_proofs));
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
    82
     log ("Number of lost proofs: " ^ string_of_int (get id num_lost_proofs));
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    83
     log ("Proof found rate: " ^
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    84
          percentage_alt (get id num_found_proofs) (get id num_lost_proofs) ^ "%");
38897
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
    85
     log ("Number of found facts: " ^ string_of_int (get id num_found_facts));
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
    86
     log ("Number of lost facts: " ^ string_of_int (get id num_lost_facts));
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    87
     log ("Fact found rate: " ^
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
    88
          percentage_alt (get id num_found_facts) (get id num_lost_facts) ^ "%"))
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    89
  else
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
    90
    ()
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    91
57154
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 55212
diff changeset
    92
val default_prover = ATP_Proof.eN (* arbitrary ATP *)
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
    93
38897
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
    94
fun with_index (i, s) = s ^ "@" ^ string_of_int i
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
    95
38896
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
    96
fun action args id ({pre, pos, log, ...} : Mirabelle.run_args) =
43710
7270ae921cf2 discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents: 43351
diff changeset
    97
  case (Position.line_of pos, Position.offset_of pos) of
7270ae921cf2 discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents: 43351
diff changeset
    98
    (SOME line_num, SOME offset) =>
7270ae921cf2 discontinued odd Position.column -- left-over from attempts at PGIP implementation;
wenzelm
parents: 43351
diff changeset
    99
    (case Prooftab.lookup (!proof_table) (line_num, offset) of
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   100
       SOME proofs =>
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   101
       let
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   102
         val thy = Proof.theory_of pre
43351
b19d95b4d736 compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
blanchet
parents: 43088
diff changeset
   103
         val {context = ctxt, facts = chained_ths, goal} = Proof.goal pre
54095
d80743f28fec simplify fudge factor code
blanchet
parents: 53551
diff changeset
   104
         val prover = AList.lookup (op =) args "prover" |> the_default default_prover
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   105
         val params as {max_facts, ...} = Sledgehammer_Commands.default_params thy args
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   106
         val default_max_facts =
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   107
           Sledgehammer_Prover_Minimize.default_max_facts_of_prover ctxt prover
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   108
         val relevance_fudge =
54095
d80743f28fec simplify fudge factor code
blanchet
parents: 53551
diff changeset
   109
           extract_relevance_fudge args Sledgehammer_MePo.default_relevance_fudge
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   110
         val subgoal = 1
52196
2281f33e8da6 redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents: 52125
diff changeset
   111
         val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal subgoal ctxt
55212
blanchet
parents: 55205
diff changeset
   112
         val ho_atp = Sledgehammer_Prover_ATP.is_ho_atp ctxt prover
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58921
diff changeset
   113
         val keywords = Thy_Header.get_keywords' ctxt
48299
5e5c6616f0fe centrally construct expensive data structures
blanchet
parents: 48293
diff changeset
   114
         val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   115
         val facts =
48289
6b65f1ad0e4b systematize lazy names in relevance filter
blanchet
parents: 48288
diff changeset
   116
           Sledgehammer_Fact.nearly_all_facts ctxt ho_atp
58921
ffdafc99f67b prefer explicit Keyword.keywords (cf. 82a71046dce8);
wenzelm
parents: 57154
diff changeset
   117
               Sledgehammer_Fact.no_fact_override keywords css_table chained_ths
48299
5e5c6616f0fe centrally construct expensive data structures
blanchet
parents: 48293
diff changeset
   118
               hyp_ts concl_t
54143
18def1c73c79 make sure add: doesn't add duplicates, and works for [no_atp] facts
blanchet
parents: 54128
diff changeset
   119
           |> Sledgehammer_Fact.drop_duplicate_facts
48406
b002cc16aa99 honor suggested MaSh weights
blanchet
parents: 48381
diff changeset
   120
           |> Sledgehammer_MePo.mepo_suggested_facts ctxt params
54095
d80743f28fec simplify fudge factor code
blanchet
parents: 53551
diff changeset
   121
                  (the_default default_max_facts max_facts) (SOME relevance_fudge) hyp_ts concl_t
51004
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 48406
diff changeset
   122
            |> map (fst o fst)
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   123
         val (found_facts, lost_facts) =
41455
wenzelm
parents: 41407
diff changeset
   124
           flat proofs |> sort_distinct string_ord
38897
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
   125
           |> map (fn fact => (find_index (curry (op =) fact) facts, fact))
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
   126
           |> List.partition (curry (op <=) 0 o fst)
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
   127
           |>> sort (prod_ord int_ord string_ord) ||> map snd
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   128
         val found_proofs = filter (forall (member (op =) facts)) proofs
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   129
         val n = length found_proofs
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   130
         val _ =
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   131
           if n = 0 then
38896
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
   132
             (add id num_failures 1; log "Failure")
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   133
           else
38896
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
   134
             (add id num_successes 1;
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
   135
              add id num_found_proofs n;
38897
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
   136
              log ("Success (" ^ string_of_int n ^ " of " ^
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
   137
                   string_of_int (length proofs) ^ " proofs)"))
38896
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
   138
         val _ = add id num_lost_proofs (length proofs - n)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
   139
         val _ = add id num_found_facts (length found_facts)
b36ab8860748 allow multiple invocations of "mirabele_sledgehammer_filter", by multiplexing on ID
blanchet
parents: 38894
diff changeset
   140
         val _ = add id num_lost_facts (length lost_facts)
38902
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   141
         val _ =
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   142
           if null found_facts then
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   143
             ()
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   144
           else
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   145
             let
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   146
               val found_weight =
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   147
                 Real.fromInt (fold (fn (n, _) => Integer.add (n * n)) found_facts 0)
38902
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   148
                   / Real.fromInt (length found_facts)
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   149
                 |> Math.sqrt |> Real.ceil
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   150
             in
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   151
               log ("Found facts (among " ^ string_of_int (length facts) ^
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   152
                    ", weight " ^ string_of_int found_weight ^ "): " ^
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   153
                    commas (map with_index found_facts))
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   154
             end
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   155
         val _ = if null lost_facts then
38897
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
   156
                   ()
92ca38d18af0 show index in fact list of all found facts
blanchet
parents: 38896
diff changeset
   157
                 else
38902
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   158
                   log ("Lost facts (among " ^ string_of_int (length facts) ^
c91be1e503bd move imperative code to where it belongs
blanchet
parents: 38900
diff changeset
   159
                        "): " ^ commas lost_facts)
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   160
       in () end
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   161
     | NONE => log "No known proof")
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   162
  | _ => ()
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   163
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   164
val proof_fileK = "proof_file"
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   165
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   166
fun invoke args =
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   167
  let
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   168
    val (pf_args, other_args) = args |> List.partition (curry (op =) proof_fileK o fst)
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   169
    val proof_file =
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   170
      (case pf_args of
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   171
        [] => error "No \"proof_file\" specified"
8450622db0c5 refactor large ML file
blanchet
parents: 55201
diff changeset
   172
      | (_, s) :: _ => s)
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   173
  in Mirabelle.register (init proof_file, action other_args, done) end
38892
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   174
eccc9e2a6412 added evaluation method for relevance filter
blanchet
parents:
diff changeset
   175
end;
38894
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   176
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   177
(* Workaround to keep the "mirabelle.pl" script happy *)
e85263e281be improve new "sledgehammer_filter" action
blanchet
parents: 38892
diff changeset
   178
structure Mirabelle_Sledgehammer_filter = Mirabelle_Sledgehammer_Filter;