src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML
author blanchet
Fri, 03 Dec 2010 18:29:14 +0100
changeset 40941 a3e6f8634a11
parent 40553 1264c9172338
child 40977 9140c5950494
permissions -rw-r--r--
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
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
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
     5
Minimization of fact list for Metis using ATPs.
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
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    13
  val minimize_facts :
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
39496
a52a4e4399c1 got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents: 39491
diff changeset
    23
open ATP_Proof
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
38021
e024504943d1 rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents: 38015
diff changeset
    26
open Sledgehammer
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
    27
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    28
(* wrapper for calling external prover *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    29
40553
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40342
diff changeset
    30
fun string_for_failure ATP_Proof.Unprovable = "Unprovable."
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40342
diff changeset
    31
  | string_for_failure ATP_Proof.TimedOut = "Timed out."
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40342
diff changeset
    32
  | string_for_failure ATP_Proof.Interrupted = "Interrupted."
1264c9172338 pick up SMT solver crashes and report them to the user/Mirabelle if desired
blanchet
parents: 40342
diff changeset
    33
  | string_for_failure _ = "Error."
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    34
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    35
fun n_facts names =
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    36
  let val n = length names in
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    37
    string_of_int n ^ " fact" ^ plural_s n ^
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
    38
    (if n > 0 then
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    39
       ": " ^ (names |> map fst
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    40
                     |> 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
    41
     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
    42
       "")
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
  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
    44
40342
3154f63e2bda don't be overly verbose in Sledgehammer's minimizer
blanchet
parents: 40341
diff changeset
    45
fun test_facts ({debug, overlord, provers, full_types, isar_proof,
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    46
                 isar_shrink_factor, ...} : params) (prover : prover)
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    47
               explicit_apply timeout i n state facts =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    48
  let
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    49
    val _ =
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    50
      Output.urgent_message ("Testing " ^ n_facts (map fst facts) ^ "...")
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
    51
    val params =
40342
3154f63e2bda don't be overly verbose in Sledgehammer's minimizer
blanchet
parents: 40341
diff changeset
    52
      {blocking = true, debug = debug, verbose = false, overlord = overlord,
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39496
diff changeset
    53
       provers = provers, full_types = full_types,
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39496
diff changeset
    54
       explicit_apply = explicit_apply, relevance_thresholds = (1.01, 1.01),
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    55
       max_relevant = NONE, isar_proof = isar_proof,
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38997
diff changeset
    56
       isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""}
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    57
    val facts =
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    58
      facts |> maps (fn (n, ths) => ths |> map (Untranslated o pair n))
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40068
diff changeset
    59
    val {goal, ...} = Proof.goal state
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
    60
    val problem =
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
    61
      {state = state, goal = goal, subgoal = i, subgoal_count = n,
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    62
       facts = facts, only = true}
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    63
    val result as {outcome, used_facts, ...} = prover params (K "") problem
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    64
  in
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 40114
diff changeset
    65
    Output.urgent_message (case outcome of
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    66
                SOME failure => string_for_failure failure
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    67
              | NONE =>
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    68
                if length used_facts = length facts then "Found proof."
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    69
                else "Found proof with " ^ n_facts used_facts ^ ".");
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
    result
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    71
  end
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    72
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    73
(* minimalization of facts *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    74
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    75
fun filter_used_facts used = filter (member (op =) used o fst)
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    76
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    77
fun sublinear_minimize _ [] p = p
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    78
  | sublinear_minimize test (x :: xs) (seen, result) =
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    79
    case test (xs @ seen) of
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    80
      result as {outcome = NONE, used_facts, ...} : prover_result =>
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    81
      sublinear_minimize test (filter_used_facts used_facts xs)
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    82
                         (filter_used_facts used_facts seen, result)
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    83
    | _ => sublinear_minimize test xs (x :: seen, result)
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    84
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
    85
(* 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
    86
   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
    87
   timeout. *)
38590
bd443b426d56 get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents: 38589
diff changeset
    88
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
    89
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    90
fun minimize_facts {provers = [], ...} _ _ _ _ = error "No prover is set."
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
    91
  | minimize_facts (params as {provers = prover_name :: _, timeout, ...}) i n
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
    92
                   state facts =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    93
  let
36378
f32c567dbcaa remove some bloat
blanchet
parents: 36375
diff changeset
    94
    val thy = Proof.theory_of state
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40553
diff changeset
    95
    val ctxt = Proof.context_of state
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40553
diff changeset
    96
    val prover = get_prover ctxt false prover_name
38590
bd443b426d56 get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents: 38589
diff changeset
    97
    val msecs = Time.toMilliseconds timeout
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 40114
diff changeset
    98
    val _ = Output.urgent_message ("Sledgehammer minimize: " ^ quote prover_name ^ ".")
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    99
    val {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
   100
    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
   101
    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
   102
      not (forall (Meson.is_fol_term thy)
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   103
                  (concl_t :: hyp_ts @ maps (map prop_of o snd) facts))
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
   104
    fun do_test timeout =
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
   105
      test_facts params prover explicit_apply timeout i n 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
   106
    val timer = Timer.startRealTimer ()
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   107
  in
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   108
    (case do_test timeout facts of
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   109
       result as {outcome = NONE, used_facts, ...} =>
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   110
       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
   111
         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
   112
         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
   113
           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
   114
           |> Time.fromMilliseconds
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   115
         val (min_thms, {message, ...}) =
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
   116
           sublinear_minimize (do_test new_timeout)
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   117
               (filter_used_facts used_facts facts) ([], result)
38094
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38093
diff changeset
   118
         val n = length min_thms
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 40114
diff changeset
   119
         val _ = Output.urgent_message (cat_lines
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   120
           ["Minimized: " ^ string_of_int n ^ " fact" ^ plural_s n] ^
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38745
diff changeset
   121
            (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
   122
               0 => ""
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
   123
             | n => " (including " ^ Int.toString n ^ " chained)") ^ ".")
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   124
       in (SOME min_thms, message) end
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   125
     | {outcome = SOME TimedOut, ...} =>
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   126
       (NONE, "Timeout: You can increase the time limit using the \"timeout\" \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   127
              \option (e.g., \"timeout = " ^
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 40205
diff changeset
   128
              string_of_int (10 + msecs div 1000) ^ "\").")
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   129
     | {outcome = SOME UnknownError, ...} =>
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   130
       (* Failure sometimes mean timeout, unfortunately. *)
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   131
       (NONE, "Failure: No proof was found with the current time limit. You \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   132
              \can increase the time limit using the \"timeout\" \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   133
              \option (e.g., \"timeout = " ^
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 40205
diff changeset
   134
              string_of_int (10 + msecs div 1000) ^ "\").")
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   135
     | {message, ...} => (NONE, "ATP error: " ^ message))
37994
b04307085a09 make TPTP generator accept full first-order formulas
blanchet
parents: 37926
diff changeset
   136
    handle ERROR msg => (NONE, "Error: " ^ msg)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   137
  end
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   138
38282
319c59682c51 move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents: 38100
diff changeset
   139
fun run_minimize params i refs state =
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   140
  let
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   141
    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
   142
    val reserved = reserved_isar_keyword_table ()
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   143
    val chained_ths = #facts (Proof.goal state)
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   144
    val facts =
40205
277508b07418 if "debug" is on, print list of relevant facts (poweruser request);
blanchet
parents: 40204
diff changeset
   145
      maps (map (apsnd single) o fact_from_ref ctxt reserved chained_ths) refs
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   146
  in
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   147
    case subgoal_count state of
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 40114
diff changeset
   148
      0 => Output.urgent_message "No subgoal!"
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   149
    | n =>
40059
6ad9081665db use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents: 39496
diff changeset
   150
      (kill_provers ();
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   151
       Output.urgent_message (#2 (minimize_facts params i n state facts)))
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   152
  end
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   153
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
   154
end;