src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML
author blanchet
Tue, 14 Sep 2010 16:34:26 +0200
changeset 39366 f58fbb959826
parent 39339 9608a5bd5d20
child 39452 70a57e40f795
permissions -rw-r--r--
handle relevance filter corner cases more gracefully; e.g. the minimizer selects 15 facts but "max_relevant = 14"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_minimize.ML
31037
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
     2
    Author:     Philipp Meyer, TU Muenchen
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
31037
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
     4
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
     5
Minimization of theorem list for Metis using automatic theorem provers.
31037
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
     6
*)
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
     7
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
     8
signature SLEDGEHAMMER_MINIMIZE =
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32510
diff changeset
     9
sig
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    10
  type locality = Sledgehammer_Filter.locality
38021
e024504943d1 rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents: 38015
diff changeset
    11
  type params = Sledgehammer.params
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
    12
38050
0511f2e62363 make Mirabelle happy
blanchet
parents: 38045
diff changeset
    13
  val minimize_theorems :
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38745
diff changeset
    14
    params -> int -> int -> Proof.state -> ((string * locality) * thm list) list
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38745
diff changeset
    15
    -> ((string * locality) * thm list) list option * string
38996
6905ba37376c generalize theorem argument parsing syntax
blanchet
parents: 38988
diff changeset
    16
  val run_minimize :
6905ba37376c generalize theorem argument parsing syntax
blanchet
parents: 38988
diff changeset
    17
    params -> int -> (Facts.ref * Attrib.src list) list -> Proof.state -> unit
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
    18
end;
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32510
diff changeset
    19
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    20
structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE =
31037
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
    21
struct
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
    22
38586
blanchet
parents: 38488
diff changeset
    23
open ATP_Systems
36142
f5e15e9aae10 make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents: 36063
diff changeset
    24
open Sledgehammer_Util
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    25
open Sledgehammer_Filter
39004
f1b465f889b5 translate the axioms to FOF once and for all ATPs
blanchet
parents: 38998
diff changeset
    26
open Sledgehammer_Translate
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    27
open Sledgehammer_Reconstruct
38021
e024504943d1 rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents: 38015
diff changeset
    28
open Sledgehammer
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
    29
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    30
(* wrapper for calling external prover *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    31
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    32
fun string_for_failure Unprovable = "Unprovable."
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    33
  | string_for_failure TimedOut = "Timed out."
39262
bdfcf2434601 better error reporting when the Sledgehammer minimizer is interrupted
blanchet
parents: 39005
diff changeset
    34
  | string_for_failure Interrupted = "Interrupted."
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    35
  | string_for_failure _ = "Unknown error."
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    36
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    37
fun n_theorems names =
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    38
  let val n = length names in
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    39
    string_of_int n ^ " theorem" ^ plural_s n ^
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    40
    (if n > 0 then
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    41
       ": " ^ (names |> map fst
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    42
                     |> sort_distinct string_ord |> space_implode " ")
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    43
     else
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    44
       "")
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    45
  end
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    46
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
    47
fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof,
38741
7635bf8918a1 get rid of "defs_relevant" feature;
blanchet
parents: 38739
diff changeset
    48
                    isar_shrink_factor, ...} : params)
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
    49
                  (prover : prover) explicit_apply timeout subgoal state
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
    50
                  axioms =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    51
  let
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    52
    val _ =
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
    53
      priority ("Testing " ^ n_theorems (map fst axioms) ^ "...")
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
    54
    val params =
38982
820b8221ed48 added "blocking" option to Sledgehammer to run in synchronous mode;
blanchet
parents: 38752
diff changeset
    55
      {blocking = true, debug = debug, verbose = verbose, overlord = overlord,
820b8221ed48 added "blocking" option to Sledgehammer to run in synchronous mode;
blanchet
parents: 38752
diff changeset
    56
       atps = atps, full_types = full_types, explicit_apply = explicit_apply,
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
    57
       relevance_thresholds = (1.01, 1.01),
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
    58
       max_relevant = SOME 65536 (* a large number *), isar_proof = isar_proof,
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
    59
       isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""}
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
    60
    val axioms = maps (fn (n, ths) => map (pair n) ths) axioms
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
    61
    val {context = ctxt, goal, ...} = Proof.goal state
39004
f1b465f889b5 translate the axioms to FOF once and for all ATPs
blanchet
parents: 38998
diff changeset
    62
    val problem =
39318
ad9a1f9b0558 implemented Auto Sledgehammer
blanchet
parents: 39262
diff changeset
    63
      {state = state, goal = goal, subgoal = subgoal,
39366
f58fbb959826 handle relevance filter corner cases more gracefully;
blanchet
parents: 39339
diff changeset
    64
       axioms = map (prepare_axiom ctxt) axioms, only = true}
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    65
    val result as {outcome, used_thm_names, ...} = prover params (K "") problem
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    66
  in
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    67
    priority (case outcome of
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    68
                NONE =>
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
    69
                if length used_thm_names = length axioms then
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    70
                  "Found proof."
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    71
                else
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    72
                  "Found proof with " ^ n_theorems used_thm_names ^ "."
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    73
              | SOME failure => string_for_failure failure);
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    74
    result
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    75
  end
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    76
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    77
(* minimalization of thms *)
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    78
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    79
fun filter_used_facts used = filter (member (op =) used o fst)
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    80
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    81
fun sublinear_minimize _ [] p = p
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    82
  | sublinear_minimize test (x :: xs) (seen, result) =
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    83
    case test (xs @ seen) of
39339
blanchet
parents: 39327
diff changeset
    84
      result as {outcome = NONE, used_thm_names, ...} : prover_result =>
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    85
      sublinear_minimize test (filter_used_facts used_thm_names xs)
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    86
                         (filter_used_facts used_thm_names seen, result)
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    87
    | _ => sublinear_minimize test xs (x :: seen, result)
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    88
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    89
(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    90
   preprocessing time is included in the estimate below but isn't part of the
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    91
   timeout. *)
38590
bd443b426d56 get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents: 38589
diff changeset
    92
val fudge_msecs = 1000
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    93
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
    94
fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set."
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
    95
  | minimize_theorems (params as {debug, atps = atp :: _, full_types,
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
    96
                                  isar_proof, isar_shrink_factor, timeout, ...})
39339
blanchet
parents: 39327
diff changeset
    97
                      i (_ : int) state axioms =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    98
  let
36378
f32c567dbcaa remove some bloat
blanchet
parents: 36375
diff changeset
    99
    val thy = Proof.theory_of state
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
   100
    val prover = get_prover_fun thy atp
38590
bd443b426d56 get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents: 38589
diff changeset
   101
    val msecs = Time.toMilliseconds timeout
bd443b426d56 get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents: 38589
diff changeset
   102
    val _ = priority ("Sledgehammer minimize: ATP " ^ quote atp ^ ".")
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
   103
    val {context = ctxt, goal, ...} = Proof.goal state
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
   104
    val (_, hyp_ts, concl_t) = strip_subgoal goal i
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
   105
    val explicit_apply =
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
   106
      not (forall (Meson.is_fol_term thy)
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
   107
                  (concl_t :: hyp_ts @ maps (map prop_of o snd) axioms))
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
   108
    fun do_test timeout =
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
   109
      test_theorems params prover explicit_apply timeout i state
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
   110
    val timer = Timer.startRealTimer ()
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   111
  in
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
   112
    (case do_test timeout axioms of
38083
c4b57f68ddb3 remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents: 38050
diff changeset
   113
       result as {outcome = NONE, pool, used_thm_names,
c4b57f68ddb3 remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents: 38050
diff changeset
   114
                  conjecture_shape, ...} =>
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   115
       let
38092
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
   116
         val time = Timer.checkRealTimer timer
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
   117
         val new_timeout =
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
   118
           Int.min (msecs, Time.toMilliseconds time + fudge_msecs)
81a003f7de0d speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents: 38084
diff changeset
   119
           |> Time.fromMilliseconds
38282
319c59682c51 move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents: 38100
diff changeset
   120
         val (min_thms, {proof, axiom_names, ...}) =
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
   121
           sublinear_minimize (do_test new_timeout)
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
   122
               (filter_used_facts used_thm_names axioms) ([], result)
38094
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38093
diff changeset
   123
         val n = length min_thms
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   124
         val _ = priority (cat_lines
38094
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38093
diff changeset
   125
           ["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38745
diff changeset
   126
            (case length (filter (curry (op =) Chained o snd o fst) min_thms) of
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
   127
               0 => ""
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
   128
             | n => " (including " ^ Int.toString n ^ " chained)") ^ ".")
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   129
       in
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   130
         (SOME min_thms,
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   131
          proof_text isar_proof
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   132
              (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
39327
blanchet
parents: 39318
diff changeset
   133
              ("Minimized command", full_types, K "", proof, axiom_names, goal,
blanchet
parents: 39318
diff changeset
   134
               i) |> fst)
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   135
       end
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   136
     | {outcome = SOME TimedOut, ...} =>
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   137
       (NONE, "Timeout: You can increase the time limit using the \"timeout\" \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   138
              \option (e.g., \"timeout = " ^
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   139
              string_of_int (10 + msecs div 1000) ^ " s\").")
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   140
     | {outcome = SOME UnknownError, ...} =>
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   141
       (* Failure sometimes mean timeout, unfortunately. *)
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   142
       (NONE, "Failure: No proof was found with the current time limit. You \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   143
              \can increase the time limit using the \"timeout\" \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   144
              \option (e.g., \"timeout = " ^
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   145
              string_of_int (10 + msecs div 1000) ^ " s\").")
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   146
     | {message, ...} => (NONE, "ATP error: " ^ message))
37994
b04307085a09 make TPTP generator accept full first-order formulas
blanchet
parents: 37926
diff changeset
   147
    handle ERROR msg => (NONE, "Error: " ^ msg)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   148
  end
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   149
38282
319c59682c51 move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents: 38100
diff changeset
   150
fun run_minimize params i refs state =
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   151
  let
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   152
    val ctxt = Proof.context_of state
38696
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38617
diff changeset
   153
    val reserved = reserved_isar_keyword_table ()
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   154
    val chained_ths = #facts (Proof.goal state)
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
   155
    val axioms =
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38745
diff changeset
   156
      maps (map (apsnd single)
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
   157
            o name_thm_pairs_from_ref ctxt reserved chained_ths) refs
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   158
  in
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   159
    case subgoal_count state of
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   160
      0 => priority "No subgoal!"
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   161
    | n =>
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38741
diff changeset
   162
      (kill_atps (); priority (#2 (minimize_theorems params i n state axioms)))
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   163
  end
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   164
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
   165
end;