src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML
author blanchet
Mon, 23 Jan 2012 17:40:32 +0100
changeset 46320 0b8b73b49848
parent 46301 e2e52c7d25c9
child 46340 cac402c486b0
permissions -rw-r--r--
renamed two files to make room for a new file
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
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    10
  type locality = ATP_Problem_Generate.locality
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    11
  type play = ATP_Proof_Reconstruct.play
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
    12
  type params = Sledgehammer_Provers.params
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
    13
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42579
diff changeset
    14
  val binary_min_facts : int Config.T
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 :
43064
b6e61d22fa61 made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
blanchet
parents: 43058
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
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
    18
    -> ((string * locality) * thm list) list option
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
    19
       * ((unit -> play) * (play -> string) * string)
38996
6905ba37376c generalize theorem argument parsing syntax
blanchet
parents: 38988
diff changeset
    20
  val run_minimize :
6905ba37376c generalize theorem argument parsing syntax
blanchet
parents: 38988
diff changeset
    21
    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
    22
end;
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32510
diff changeset
    23
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    24
structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE =
31037
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
    25
struct
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
    26
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43064
diff changeset
    27
open ATP_Util
39496
a52a4e4399c1 got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents: 39491
diff changeset
    28
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    29
open ATP_Problem_Generate
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    30
open ATP_Proof_Reconstruct
36142
f5e15e9aae10 make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents: 36063
diff changeset
    31
open Sledgehammer_Util
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    32
open Sledgehammer_Filter
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
    33
open Sledgehammer_Provers
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
    34
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    35
(* wrapper for calling external prover *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    36
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    37
fun n_facts names =
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    38
  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
    39
    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
    40
    (if n > 0 then
45366
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
    41
       ": " ^ (names |> map fst |> sort_distinct string_ord
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
    42
                     |> 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
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
    47
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
    48
42724
4d6bcf846759 added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
blanchet
parents: 42646
diff changeset
    49
fun test_facts ({debug, verbose, overlord, provers, max_mono_iters,
46301
e2e52c7d25c9 renamed "sound" option to "strict"
blanchet
parents: 45707
diff changeset
    50
                 max_new_mono_instances, type_enc, strict, lam_trans,
e2e52c7d25c9 renamed "sound" option to "strict"
blanchet
parents: 45707
diff changeset
    51
                 isar_proof, isar_shrink_factor, preplay_timeout, ...} : params)
43064
b6e61d22fa61 made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
blanchet
parents: 43058
diff changeset
    52
               silent (prover : prover) 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 ^ ")"
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43064
diff changeset
    58
           else "") ^ "...")
41742
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
    59
    val {goal, ...} = Proof.goal state
44463
c471a2b48fa1 make sure that all facts are passed to ATP from minimizer
blanchet
parents: 43630
diff changeset
    60
    val facts =
c471a2b48fa1 make sure that all facts are passed to ATP from minimizer
blanchet
parents: 43630
diff changeset
    61
      facts |> maps (fn (n, ths) => ths |> map (Untranslated_Fact o pair n))
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
    62
    val params =
42060
889d767ce5f4 make Minimizer honor "verbose" and "debug" options better
blanchet
parents: 41824
diff changeset
    63
      {debug = debug, verbose = verbose, overlord = overlord, blocking = true,
46301
e2e52c7d25c9 renamed "sound" option to "strict"
blanchet
parents: 45707
diff changeset
    64
       provers = provers, type_enc = type_enc, strict = strict,
45514
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45372
diff changeset
    65
       lam_trans = lam_trans, relevance_thresholds = (1.01, 1.01),
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45372
diff changeset
    66
       max_relevant = SOME (length facts), max_mono_iters = max_mono_iters,
42740
31334a7b109d renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
blanchet
parents: 42724
diff changeset
    67
       max_new_mono_instances = max_new_mono_instances, isar_proof = isar_proof,
45706
418846ea4f99 renamed "slicing" to "slice"
blanchet
parents: 45574
diff changeset
    68
       isar_shrink_factor = isar_shrink_factor, slice = false,
45707
6bf7eec9b153 added "minimize" option for more control over automatic minimization
blanchet
parents: 45706
diff changeset
    69
       minimize = SOME false, timeout = timeout,
6bf7eec9b153 added "minimize" option for more control over automatic minimization
blanchet
parents: 45706
diff changeset
    70
       preplay_timeout = preplay_timeout, expect = ""}
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
    71
    val problem =
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
    72
      {state = state, goal = goal, subgoal = i, subgoal_count = n,
41741
839d1488045f renamed field
blanchet
parents: 41491
diff changeset
    73
       facts = facts, smt_filter = NONE}
45370
bab52dafa63a use "Time.time" rather than milliseconds internally
blanchet
parents: 45369
diff changeset
    74
    val result as {outcome, used_facts, run_time, ...} =
45520
2b1dde0b1c30 thread in additional options to minimizer
blanchet
parents: 45519
diff changeset
    75
      prover params (K (K (K ""))) problem
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    76
  in
43259
30c141dc22d6 killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
blanchet
parents: 43166
diff changeset
    77
    print silent
30c141dc22d6 killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
blanchet
parents: 43166
diff changeset
    78
          (fn () =>
30c141dc22d6 killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
blanchet
parents: 43166
diff changeset
    79
              case outcome of
30c141dc22d6 killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
blanchet
parents: 43166
diff changeset
    80
                SOME failure => string_for_failure failure
30c141dc22d6 killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
blanchet
parents: 43166
diff changeset
    81
              | NONE =>
30c141dc22d6 killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
blanchet
parents: 43166
diff changeset
    82
                "Found proof" ^
30c141dc22d6 killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
blanchet
parents: 43166
diff changeset
    83
                 (if length used_facts = length facts then ""
45370
bab52dafa63a use "Time.time" rather than milliseconds internally
blanchet
parents: 45369
diff changeset
    84
                  else " with " ^ n_facts used_facts) ^
bab52dafa63a use "Time.time" rather than milliseconds internally
blanchet
parents: 45369
diff changeset
    85
                 " (" ^ string_from_time run_time ^ ").");
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
    86
    result
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    87
  end
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    88
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
    89
(* minimalization of facts *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    90
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    91
(* Give the external prover some slack. The ATP gets further slack because the
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    92
   Sledgehammer preprocessing time is included in the estimate below but isn't
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    93
   part of the timeout. *)
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    94
val slack_msecs = 200
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    95
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    96
fun new_timeout timeout run_time =
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    97
  Int.min (Time.toMilliseconds timeout,
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    98
           Time.toMilliseconds run_time + slack_msecs)
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
    99
  |> Time.fromMilliseconds
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   100
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   101
(* The linear algorithm usually outperforms the binary algorithm when over 60%
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   102
   of the facts are actually needed. The binary algorithm is much more
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   103
   appropriate for provers that cannot return the list of used facts and hence
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   104
   returns all facts as used. Since we cannot know in advance how many facts are
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   105
   actually needed, we heuristically set the threshold to 10 facts. *)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42579
diff changeset
   106
val binary_min_facts =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42579
diff changeset
   107
  Attrib.setup_config_int @{binding sledgehammer_minimize_binary_min_facts}
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   108
                          (K 20)
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   109
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   110
fun linear_minimize test timeout result xs =
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   111
  let
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   112
    fun min _ [] p = p
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   113
      | min timeout (x :: xs) (seen, result) =
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   114
        case test timeout (xs @ seen) of
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   115
          result as {outcome = NONE, used_facts, run_time, ...}
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   116
          : prover_result =>
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   117
          min (new_timeout timeout run_time) (filter_used_facts used_facts xs)
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   118
              (filter_used_facts used_facts seen, result)
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   119
        | _ => min timeout xs (x :: seen, result)
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   120
  in min timeout xs ([], result) end
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   121
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   122
fun binary_minimize test timeout result xs =
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   123
  let
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   124
    fun min depth (result as {run_time, ...} : prover_result) sup
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   125
            (xs as _ :: _ :: _) =
41743
blanchet
parents: 41742
diff changeset
   126
        let
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   127
          val (l0, r0) = chop (length xs div 2) xs
41743
blanchet
parents: 41742
diff changeset
   128
(*
45366
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
   129
          val _ = warning (replicate_string depth " " ^ "{ " ^
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
   130
                           "sup: " ^ n_facts (map fst sup))
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
   131
          val _ = warning (replicate_string depth " " ^ "  " ^
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
   132
                           "xs: " ^ n_facts (map fst xs))
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
   133
          val _ = warning (replicate_string depth " " ^ "  " ^
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
   134
                           "l0: " ^ n_facts (map fst l0))
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
   135
          val _ = warning (replicate_string depth " " ^ "  " ^
4339763edd55 speed up binary minimizer in practice by preferring the first half of the used facts (which are likelier to be relevant) to the second half
blanchet
parents: 44463
diff changeset
   136
                           "r0: " ^ n_facts (map fst r0))
41743
blanchet
parents: 41742
diff changeset
   137
*)
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   138
          val depth = depth + 1
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   139
          val timeout = new_timeout timeout run_time
41743
blanchet
parents: 41742
diff changeset
   140
        in
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   141
          case test timeout (sup @ l0) of
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   142
            result as {outcome = NONE, used_facts, ...} =>
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   143
            min depth result (filter_used_facts used_facts sup)
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   144
                      (filter_used_facts used_facts l0)
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   145
          | _ =>
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   146
            case test timeout (sup @ r0) of
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   147
              result as {outcome = NONE, used_facts, ...} =>
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   148
              min depth result (filter_used_facts used_facts sup)
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   149
                        (filter_used_facts used_facts r0)
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   150
            | _ =>
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   151
              let
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   152
                val (sup_r0, (l, result)) = min depth result (sup @ r0) l0
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   153
                val (sup, r0) =
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   154
                  (sup, r0) |> pairself (filter_used_facts (map fst sup_r0))
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   155
                val (sup_l, (r, result)) = min depth result (sup @ l) r0
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   156
                val sup = sup |> filter_used_facts (map fst sup_l)
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   157
              in (sup, (l @ r, result)) end
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   158
        end
41743
blanchet
parents: 41742
diff changeset
   159
(*
blanchet
parents: 41742
diff changeset
   160
        |> tap (fn _ => warning (replicate_string depth " " ^ "}"))
blanchet
parents: 41742
diff changeset
   161
*)
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   162
      | min _ result sup xs = (sup, (xs, result))
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   163
  in
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   164
    case snd (min 0 result [] xs) of
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   165
      ([x], result as {run_time, ...}) =>
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   166
      (case test (new_timeout timeout run_time) [] of
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   167
         result as {outcome = NONE, ...} => ([], result)
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   168
       | _ => ([x], result))
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   169
    | p => p
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   170
  end
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   171
43064
b6e61d22fa61 made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
blanchet
parents: 43058
diff changeset
   172
fun minimize_facts prover_name (params as {timeout, ...}) silent i n state
b6e61d22fa61 made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
blanchet
parents: 43058
diff changeset
   173
                   facts =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   174
  let
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40553
diff changeset
   175
    val ctxt = Proof.context_of state
45574
7a39df11bcf6 be more silent when auto minimizing
blanchet
parents: 45520
diff changeset
   176
    val prover =
7a39df11bcf6 be more silent when auto minimizing
blanchet
parents: 45520
diff changeset
   177
      get_prover ctxt (if silent then Auto_Minimize else Minimize) prover_name
43058
5f8bac7a2945 minimize automatically even if Metis failed, if the external prover was really fast
blanchet
parents: 43052
diff changeset
   178
    val _ = print silent (fn () => "Sledgehammer minimizer: " ^
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   179
                                   quote prover_name ^ ".")
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   180
    fun test timeout = test_facts params silent prover timeout i n state
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   181
  in
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   182
    (case test timeout facts of
45371
blanchet
parents: 45370
diff changeset
   183
       result as {outcome = NONE, used_facts, run_time, ...} =>
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   184
       let
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   185
         val facts = filter_used_facts used_facts facts
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   186
         val min =
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42579
diff changeset
   187
           if length facts >= Config.get ctxt binary_min_facts then
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   188
             binary_minimize
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   189
           else
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   190
             linear_minimize
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   191
         val (min_facts, {preplay, message, message_tail, ...}) =
45371
blanchet
parents: 45370
diff changeset
   192
           min test (new_timeout timeout run_time) result facts
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   193
         val _ = print silent (fn () => cat_lines
43630
e42ccb132305 made minimizer informative output accurate
blanchet
parents: 43626
diff changeset
   194
           ["Minimized to " ^ n_facts (map fst min_facts)] ^
43577
78c2f14b35df clarify minimizer output
blanchet
parents: 43572
diff changeset
   195
            (case length (filter (curry (op =) Chained o snd o fst) min_facts) of
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
   196
               0 => ""
43577
78c2f14b35df clarify minimizer output
blanchet
parents: 43572
diff changeset
   197
             | n => "\n(including " ^ string_of_int n ^ " chained)") ^ ".")
78c2f14b35df clarify minimizer output
blanchet
parents: 43572
diff changeset
   198
       in (SOME min_facts, (preplay, message, message_tail)) end
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   199
     | {outcome = SOME TimedOut, preplay, ...} =>
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   200
       (NONE,
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   201
        (preplay,
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   202
         fn _ => "Timeout: You can increase the time limit using the \
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   203
                 \\"timeout\" option (e.g., \"timeout = " ^
45371
blanchet
parents: 45370
diff changeset
   204
                 string_of_int (10 + Time.toMilliseconds timeout div 1000) ^
blanchet
parents: 45370
diff changeset
   205
                 "\").",
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
   206
         ""))
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   207
     | {preplay, message, ...} =>
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
   208
       (NONE, (preplay, prefix "Prover error: " o message, "")))
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43085
diff changeset
   209
    handle ERROR msg =>
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   210
           (NONE, (K (Failed_to_Play plain_metis), fn _ => "Error: " ^ msg, ""))
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   211
  end
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   212
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
   213
fun run_minimize (params as {provers, ...}) i refs state =
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   214
  let
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   215
    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
   216
    val reserved = reserved_isar_keyword_table ()
43043
1406f6fc5dc3 normalize indices in chained facts to make sure that backtick facts (which often result in different names) are recognized + changed definition of urgent messages
blanchet
parents: 43033
diff changeset
   217
    val chained_ths = normalize_chained_theorems (#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
   218
    val facts =
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   219
      refs
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
   220
      |> maps (map (apsnd single) o fact_from_ref ctxt reserved chained_ths)
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   221
  in
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   222
    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
   223
      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
   224
    | 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
   225
             [] => 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
   226
           | 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
   227
             (kill_provers ();
43064
b6e61d22fa61 made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
blanchet
parents: 43058
diff changeset
   228
              minimize_facts prover params false i n state facts
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
   229
              |> (fn (_, (preplay, message, message_tail)) =>
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
   230
                     message (preplay ()) ^ message_tail
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
   231
                     |> Output.urgent_message))
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   232
  end
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   233
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
   234
end;