src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML
author wenzelm
Sun, 30 Jan 2011 13:02:18 +0100
changeset 41648 6d736d983d5c
parent 41491 a2ad5b824051
child 41741 839d1488045f
permissions -rw-r--r--
clarified example settings for Proof General;
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
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
     5
Minimization of fact list for Metis using external 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
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
    11
  type params = Sledgehammer_Provers.params
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
    12
41335
66edbd0f7a2e added "smt_triggers" option to experiment with triggers in Sledgehammer;
blanchet
parents: 41277
diff changeset
    13
  val binary_min_facts : int Unsynchronized.ref
41255
a80024d7b71b added debugging option to find out how good the relevance filter was at identifying relevant facts
blanchet
parents: 41242
diff changeset
    14
  val filter_used_facts : ''a list -> (''a * 'b) list -> (''a * 'b) list
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    15
  val minimize_facts :
41265
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
    16
    string -> params -> bool -> int -> int -> Proof.state
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
    17
    -> ((string * locality) * thm list) list
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38745
diff changeset
    18
    -> ((string * locality) * thm list) list option * string
38996
6905ba37376c generalize theorem argument parsing syntax
blanchet
parents: 38988
diff changeset
    19
  val run_minimize :
6905ba37376c generalize theorem argument parsing syntax
blanchet
parents: 38988
diff changeset
    20
    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
    21
end;
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32510
diff changeset
    22
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    23
structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE =
31037
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
    24
struct
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
    25
39496
a52a4e4399c1 got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents: 39491
diff changeset
    26
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
    27
open Sledgehammer_Util
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    28
open Sledgehammer_Filter
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
    29
open Sledgehammer_Provers
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
    30
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    31
(* wrapper for calling external prover *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    32
41277
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    33
fun short_string_for_failure ATP_Proof.Unprovable = "Unprovable."
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    34
  | short_string_for_failure ATP_Proof.TimedOut = "Timed out."
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    35
  | short_string_for_failure ATP_Proof.Interrupted = "Interrupted."
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    36
  | short_string_for_failure _ = "Error."
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    37
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    38
fun n_facts names =
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    39
  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
    40
    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
    41
    (if n > 0 then
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    42
       ": " ^ (names |> map fst
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    43
                     |> 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
    44
     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
    45
       "")
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
  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
    47
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
    48
fun print silent f = if silent then () else Output.urgent_message (f ())
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
    49
41277
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    50
fun test_facts ({debug, verbose, overlord, provers, type_sys, isar_proof,
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
    51
                 isar_shrink_factor, ...} : params) silent (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
    52
               explicit_apply timeout i n state facts =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    53
  let
41277
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    54
    val _ =
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    55
      print silent (fn () =>
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    56
          "Testing " ^ n_facts (map fst facts) ^
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    57
          (if verbose then " (timeout: " ^ string_from_time timeout ^ ")"
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    58
          else "") ^ "...")
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
    59
    val params =
41208
1b28c43a7074 make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
blanchet
parents: 41138
diff changeset
    60
      {debug = debug, verbose = false, overlord = overlord, blocking = true,
41138
eb80538166b6 implemented partially-typed "tags" type encoding
blanchet
parents: 41134
diff changeset
    61
       provers = provers, type_sys = type_sys, explicit_apply = explicit_apply,
eb80538166b6 implemented partially-typed "tags" type encoding
blanchet
parents: 41134
diff changeset
    62
       relevance_thresholds = (1.01, 1.01), max_relevant = NONE,
eb80538166b6 implemented partially-typed "tags" type encoding
blanchet
parents: 41134
diff changeset
    63
       isar_proof = isar_proof, isar_shrink_factor = isar_shrink_factor,
eb80538166b6 implemented partially-typed "tags" type encoding
blanchet
parents: 41134
diff changeset
    64
       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
    65
    val facts =
41090
b98fe4de1ecd renamings
blanchet
parents: 41087
diff changeset
    66
      facts |> maps (fn (n, ths) => ths |> map (Untranslated_Fact o pair n))
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40068
diff changeset
    67
    val {goal, ...} = Proof.goal state
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
    68
    val problem =
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
    69
      {state = state, goal = goal, subgoal = i, subgoal_count = n,
41242
8edeb1dbbc76 run the SMT relevance filter only once, then run the normalization/monomorphization code once _per class_ of SMT solvers
blanchet
parents: 41208
diff changeset
    70
       facts = facts, smt_head = 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
    71
    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
    72
  in
41277
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    73
    print silent (fn () =>
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    74
        case outcome of
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    75
          SOME failure => short_string_for_failure failure
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    76
        | NONE =>
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    77
          if length used_facts = length facts then "Found proof."
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    78
          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
    79
    result
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    80
  end
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    81
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
    82
(* minimalization of facts *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    83
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
    84
(* The sublinear algorithm works well in almost all situations, except when the
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
    85
   external prover cannot return the list of used facts and hence returns all
41267
958fee9ec275 lower threshold where the binary algorithm kick in and use the same value for automatic minimization
blanchet
parents: 41265
diff changeset
    86
   facts as used. In that case, the binary algorithm is much more appropriate.
958fee9ec275 lower threshold where the binary algorithm kick in and use the same value for automatic minimization
blanchet
parents: 41265
diff changeset
    87
   We can usually detect the situation by looking at the number of used facts
958fee9ec275 lower threshold where the binary algorithm kick in and use the same value for automatic minimization
blanchet
parents: 41265
diff changeset
    88
   reported by the prover. *)
41335
66edbd0f7a2e added "smt_triggers" option to experiment with triggers in Sledgehammer;
blanchet
parents: 41277
diff changeset
    89
val binary_min_facts = Unsynchronized.ref 20
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
    90
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
    91
fun filter_used_facts used = filter (member (op =) used o fst)
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    92
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    93
fun sublinear_minimize _ [] p = p
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    94
  | sublinear_minimize test (x :: xs) (seen, result) =
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    95
    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
    96
      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
    97
      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
    98
                         (filter_used_facts used_facts seen, result)
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
    99
    | _ => sublinear_minimize test xs (x :: seen, result)
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   100
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   101
fun binary_minimize test xs =
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   102
  let
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   103
    fun p xs = #outcome (test xs : prover_result) = NONE
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   104
    fun split [] p = p
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   105
      | split [h] (l, r) = (h :: l, r)
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   106
      | split (h1 :: h2 :: t) (l, r) = split t (h1 :: l, h2 :: r)
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   107
    fun min _ [] = raise Empty
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   108
      | min _ [s0] = [s0]
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   109
      | min sup xs =
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   110
        let val (l0, r0) = split xs ([], []) in
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   111
          if p (sup @ l0) then
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   112
            min sup l0
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   113
          else if p (sup @ r0) then
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   114
            min sup r0
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   115
          else
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   116
            let
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   117
              val l = min (sup @ r0) l0
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   118
              val r = min (sup @ l) r0
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   119
            in l @ r end
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   120
        end
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   121
    val xs =
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   122
      case min [] xs of
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   123
        [x] => if p [] then [] else [x]
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   124
      | xs => xs
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   125
  in (xs, test xs) end
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   126
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   127
(* Give the external prover some slack. The ATP gets further slack because the
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   128
   Sledgehammer preprocessing time is included in the estimate below but isn't
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   129
   part of the timeout. *)
41277
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
   130
val slack_msecs = 200
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
   131
41265
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   132
fun minimize_facts prover_name (params as {timeout, ...}) silent i n state
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   133
                   facts =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   134
  let
36378
f32c567dbcaa remove some bloat
blanchet
parents: 36375
diff changeset
   135
    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
   136
    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
   137
    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
   138
    val msecs = Time.toMilliseconds timeout
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   139
    val _ = print silent (fn () => "Sledgehammer minimize: " ^
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   140
                                   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
   141
    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
   142
    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
   143
    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
   144
      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
   145
                  (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
   146
    fun do_test timeout =
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   147
      test_facts params silent 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
   148
    val timer = Timer.startRealTimer ()
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   149
  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
   150
    (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
   151
       result as {outcome = NONE, used_facts, ...} =>
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   152
       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
   153
         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
   154
         val new_timeout =
41277
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
   155
           Int.min (msecs, Time.toMilliseconds time + slack_msecs)
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
   156
           |> Time.fromMilliseconds
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   157
         val facts = filter_used_facts used_facts facts
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   158
         val (min_thms, {message, ...}) =
41335
66edbd0f7a2e added "smt_triggers" option to experiment with triggers in Sledgehammer;
blanchet
parents: 41277
diff changeset
   159
           if length facts >= !binary_min_facts then
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   160
             binary_minimize (do_test new_timeout) facts
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   161
           else
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   162
             sublinear_minimize (do_test new_timeout) facts ([], result)
38094
d01b8119b2e0 better error and minimizer output
blanchet
parents: 38093
diff changeset
   163
         val n = length min_thms
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   164
         val _ = print silent (fn () => cat_lines
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   165
           ["Minimized: " ^ string_of_int n ^ " fact" ^ plural_s n] ^
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38745
diff changeset
   166
            (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
   167
               0 => ""
41491
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41335
diff changeset
   168
             | n => " (including " ^ string_of_int n ^ " chained)") ^ ".")
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   169
       in (SOME min_thms, message) end
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   170
     | {outcome = SOME TimedOut, ...} =>
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   171
       (NONE, "Timeout: You can increase the time limit using the \"timeout\" \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   172
              \option (e.g., \"timeout = " ^
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 40205
diff changeset
   173
              string_of_int (10 + msecs div 1000) ^ "\").")
41259
13972ced98d9 more precise error messages in "verbose" (or "debug") mode, following this morning's permission debacle
blanchet
parents: 41255
diff changeset
   174
     | {outcome = SOME (UnknownError _), ...} =>
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   175
       (* Failure sometimes mean timeout, unfortunately. *)
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   176
       (NONE, "Failure: No proof was found with the current time limit. You \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   177
              \can increase the time limit using the \"timeout\" \
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   178
              \option (e.g., \"timeout = " ^
40341
03156257040f standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents: 40205
diff changeset
   179
              string_of_int (10 + msecs div 1000) ^ "\").")
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   180
     | {message, ...} => (NONE, "Prover error: " ^ message))
37994
b04307085a09 make TPTP generator accept full first-order formulas
blanchet
parents: 37926
diff changeset
   181
    handle ERROR msg => (NONE, "Error: " ^ msg)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   182
  end
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   183
41265
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   184
fun run_minimize (params as {provers, ...}) i refs state =
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   185
  let
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   186
    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
   187
    val reserved = reserved_isar_keyword_table ()
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   188
    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
   189
    val facts =
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   190
      refs
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   191
      |> maps (map (apsnd single) o fact_from_ref ctxt reserved chained_ths)
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   192
  in
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   193
    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
   194
      0 => Output.urgent_message "No subgoal!"
41265
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   195
    | n => case provers of
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   196
             [] => error "No prover is set."
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   197
           | prover :: _ =>
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   198
             (kill_provers ();
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   199
              minimize_facts prover params false i n state facts
a393d6d8e198 let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
blanchet
parents: 41259
diff changeset
   200
              |> #2 |> Output.urgent_message)
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   201
  end
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   202
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
   203
end;