src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML
author wenzelm
Sat, 09 Nov 2013 18:00:36 +0100
changeset 54381 9c1f21365326
parent 53800 ac1ec5065316
child 54119 2c13cb4a057d
permissions -rw-r--r--
tuned;
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
46340
cac402c486b0 separate orthogonal components
blanchet
parents: 46320
diff changeset
    10
  type stature = ATP_Problem_Generate.stature
52555
6811291d1869 moved code -> easier debugging
smolkas
parents: 52031
diff changeset
    11
  type play = Sledgehammer_Reconstructor.play
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
    12
  type mode = Sledgehammer_Provers.mode
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
    13
  type params = Sledgehammer_Provers.params
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
    14
  type prover = Sledgehammer_Provers.prover
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
    15
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42579
diff changeset
    16
  val binary_min_facts : int Config.T
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
    17
  val auto_minimize_min_facts : int Config.T
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
    18
  val auto_minimize_max_time : real Config.T
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    19
  val minimize_facts :
48399
4bacc8983b3d learn from SMT proofs when they can be minimized by Metis
blanchet
parents: 48396
diff changeset
    20
    (string -> thm list -> unit) -> string -> params -> bool -> int -> int
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
    21
    -> Proof.state -> play Lazy.lazy option
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
    22
    -> ((string * stature) * thm list) list
46340
cac402c486b0 separate orthogonal components
blanchet
parents: 46320
diff changeset
    23
    -> ((string * stature) * thm list) list option
50669
84c7cf36b2e0 use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents: 50668
diff changeset
    24
       * (play Lazy.lazy * (play -> string) * string)
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
    25
  val get_minimizing_prover :
48399
4bacc8983b3d learn from SMT proofs when they can be minimized by Metis
blanchet
parents: 48396
diff changeset
    26
    Proof.context -> mode -> (string -> thm list -> unit) -> string -> prover
38996
6905ba37376c generalize theorem argument parsing syntax
blanchet
parents: 38988
diff changeset
    27
  val run_minimize :
48399
4bacc8983b3d learn from SMT proofs when they can be minimized by Metis
blanchet
parents: 48396
diff changeset
    28
    params -> (string -> thm list -> unit) -> int
4bacc8983b3d learn from SMT proofs when they can be minimized by Metis
blanchet
parents: 48396
diff changeset
    29
    -> (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
    30
end;
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32510
diff changeset
    31
38988
483879af0643 finished renaming
blanchet
parents: 38986
diff changeset
    32
structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE =
31037
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
    33
struct
ac8669134e7a added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff changeset
    34
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43064
diff changeset
    35
open ATP_Util
39496
a52a4e4399c1 got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents: 39491
diff changeset
    36
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    37
open ATP_Problem_Generate
51200
260cb10aac4b minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents: 51191
diff changeset
    38
open ATP_Systems
36142
f5e15e9aae10 make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents: 36063
diff changeset
    39
open Sledgehammer_Util
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
    40
open Sledgehammer_Fact
52555
6811291d1869 moved code -> easier debugging
smolkas
parents: 52031
diff changeset
    41
open Sledgehammer_Reconstructor
49914
23e36a4d28f1 refactor code
blanchet
parents: 48799
diff changeset
    42
open Sledgehammer_Reconstruct
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
    43
open Sledgehammer_Provers
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
    44
36370
a4f601daa175 centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents: 36369
diff changeset
    45
(* wrapper for calling external prover *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    46
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    47
fun n_facts names =
38698
d19c3a7ce38b clean handling of whether a fact is chained or not;
blanchet
parents: 38696
diff changeset
    48
  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
    49
    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
    50
    (if n > 0 then
48085
ff5e900d7b1a avoid dumping definitions several times in LEO-II proofs
blanchet
parents: 47531
diff changeset
    51
       ": " ^ (names |> map fst |> sort 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
    52
     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
    53
       "")
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
    54
  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
    55
41091
0afdf5cde874 implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents: 41090
diff changeset
    56
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
    57
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    58
fun test_facts ({debug, verbose, overlord, spy, provers, max_mono_iters,
46301
e2e52c7d25c9 renamed "sound" option to "strict"
blanchet
parents: 45707
diff changeset
    59
                 max_new_mono_instances, type_enc, strict, lam_trans,
53764
eba0d1c069b8 merged "isar_try0" and "isar_minimize" options
blanchet
parents: 53763
diff changeset
    60
                 uncurried_aliases, isar_proofs, isar_compress, isar_try0,
eba0d1c069b8 merged "isar_try0" and "isar_minimize" options
blanchet
parents: 53763
diff changeset
    61
                 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
    62
               silent (prover : prover) timeout i n state facts =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
    63
  let
41277
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    64
    val _ =
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    65
      print silent (fn () =>
1369c27c6966 reduce the minimizer slack and add verbose information
blanchet
parents: 41267
diff changeset
    66
          "Testing " ^ n_facts (map fst facts) ^
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
    67
          (if verbose then
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
    68
             case timeout of
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
    69
               SOME timeout => " (timeout: " ^ string_of_time timeout ^ ")"
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
    70
             | _ => ""
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
    71
           else
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
    72
             "") ^ "...")
41742
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
    73
    val {goal, ...} = Proof.goal state
51005
ce4290c33d73 eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents: 50669
diff changeset
    74
    val facts = facts |> maps (fn (n, ths) => map (pair n) ths)
38100
e458a0dd3dc1 use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents: 38094
diff changeset
    75
    val params =
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    76
      {debug = debug, verbose = verbose, overlord = overlord, spy = spy, blocking = true,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    77
       provers = provers, type_enc = type_enc, strict = strict, lam_trans = lam_trans,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    78
       uncurried_aliases = uncurried_aliases, learn = false, fact_filter = NONE,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    79
       max_facts = SOME (length facts), fact_thresholds = (1.01, 1.01),
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    80
       max_mono_iters = max_mono_iters, max_new_mono_instances = max_new_mono_instances,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    81
       isar_proofs = isar_proofs, isar_compress = isar_compress, isar_try0 = isar_try0,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    82
       slice = false, minimize = SOME false, timeout = timeout, preplay_timeout = preplay_timeout,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    83
       expect = ""}
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
    84
    val problem =
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    85
      {state = state, goal = goal, subgoal = i, subgoal_count = n, factss = [("", facts)]}
45370
bab52dafa63a use "Time.time" rather than milliseconds internally
blanchet
parents: 45369
diff changeset
    86
    val result as {outcome, used_facts, run_time, ...} =
45520
2b1dde0b1c30 thread in additional options to minimizer
blanchet
parents: 45519
diff changeset
    87
      prover params (K (K (K ""))) problem
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    88
  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
    89
    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
    90
          (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
    91
              case outcome of
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52632
diff changeset
    92
                SOME failure => string_of_atp_failure failure
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
    93
              | 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
    94
                "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
    95
                 (if length used_facts = length facts then ""
45370
bab52dafa63a use "Time.time" rather than milliseconds internally
blanchet
parents: 45369
diff changeset
    96
                  else " with " ^ n_facts used_facts) ^
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
    97
                 " (" ^ string_of_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
    98
    result
36223
217ca1273786 make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents: 36143
diff changeset
    99
  end
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   100
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
   101
(* minimalization of facts *)
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   102
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   103
(* 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
   104
   Sledgehammer preprocessing time is included in the estimate below but isn't
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   105
   part of the timeout. *)
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   106
val slack_msecs = 200
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   107
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   108
fun new_timeout NONE _ = NONE
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   109
  | new_timeout (SOME timeout) run_time =
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   110
    Int.min (Time.toMilliseconds timeout,
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   111
             Time.toMilliseconds run_time + slack_msecs)
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   112
    |> Time.fromMilliseconds |> SOME
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   113
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   114
(* The linear algorithm usually outperforms the binary algorithm when over 60%
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   115
   of the facts are actually needed. The binary algorithm is much more
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   116
   appropriate for provers that cannot return the list of used facts and hence
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   117
   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
   118
   actually needed, we heuristically set the threshold to 10 facts. *)
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42579
diff changeset
   119
val binary_min_facts =
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42579
diff changeset
   120
  Attrib.setup_config_int @{binding sledgehammer_minimize_binary_min_facts}
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   121
                          (K 20)
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   122
val auto_minimize_min_facts =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   123
  Attrib.setup_config_int @{binding sledgehammer_auto_minimize_min_facts}
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   124
      (fn generic => Config.get_generic generic binary_min_facts)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   125
val auto_minimize_max_time =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   126
  Attrib.setup_config_real @{binding sledgehammer_auto_minimize_max_time}
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   127
                           (K 5.0)
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   128
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   129
fun linear_minimize test timeout result xs =
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   130
  let
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   131
    fun min _ [] p = p
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   132
      | min timeout (x :: xs) (seen, result) =
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   133
        case test timeout (xs @ seen) of
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   134
          result as {outcome = NONE, used_facts, run_time, ...}
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   135
          : prover_result =>
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   136
          min (new_timeout timeout run_time)
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   137
              (filter_used_facts true used_facts xs)
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   138
              (filter_used_facts false used_facts seen, result)
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   139
        | _ => min timeout xs (x :: seen, result)
48799
5c9356f8c968 warn users about unused "using" facts
blanchet
parents: 48798
diff changeset
   140
  in min timeout xs ([], result) end
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   141
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   142
fun binary_minimize test timeout result xs =
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   143
  let
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   144
    fun min depth (result as {run_time, ...} : prover_result) sup
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   145
            (xs as _ :: _ :: _) =
41743
blanchet
parents: 41742
diff changeset
   146
        let
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   147
          val (l0, r0) = chop (length xs div 2) xs
41743
blanchet
parents: 41742
diff changeset
   148
(*
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
   149
          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
   150
                           "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
   151
          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
   152
                           "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
   153
          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
   154
                           "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
   155
          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
   156
                           "r0: " ^ n_facts (map fst r0))
41743
blanchet
parents: 41742
diff changeset
   157
*)
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   158
          val depth = depth + 1
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   159
          val timeout = new_timeout timeout run_time
41743
blanchet
parents: 41742
diff changeset
   160
        in
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   161
          case test timeout (sup @ l0) of
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   162
            result as {outcome = NONE, used_facts, ...} =>
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   163
            min depth result (filter_used_facts true used_facts sup)
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   164
                      (filter_used_facts true used_facts l0)
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   165
          | _ =>
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   166
            case test timeout (sup @ r0) of
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   167
              result as {outcome = NONE, used_facts, ...} =>
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   168
              min depth result (filter_used_facts true used_facts sup)
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   169
                        (filter_used_facts true used_facts r0)
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   170
            | _ =>
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   171
              let
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   172
                val (sup_r0, (l, result)) = min depth result (sup @ r0) l0
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   173
                val (sup, r0) =
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   174
                  (sup, r0)
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   175
                  |> pairself (filter_used_facts true (map fst sup_r0))
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   176
                val (sup_l, (r, result)) = min depth result (sup @ l) r0
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   177
                val sup = sup |> filter_used_facts true (map fst sup_l)
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   178
              in (sup, (l @ r, result)) end
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   179
        end
41743
blanchet
parents: 41742
diff changeset
   180
(*
blanchet
parents: 41742
diff changeset
   181
        |> tap (fn _ => warning (replicate_string depth " " ^ "}"))
blanchet
parents: 41742
diff changeset
   182
*)
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   183
      | min _ result sup xs = (sup, (xs, result))
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   184
  in
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   185
    case snd (min 0 result [] xs) of
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   186
      ([x], result as {run_time, ...}) =>
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   187
      (case test (new_timeout timeout run_time) [] of
45367
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   188
         result as {outcome = NONE, ...} => ([], result)
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   189
       | _ => ([x], result))
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   190
    | p => p
cb54f1b34cf9 shortcut binary minimization algorithm
blanchet
parents: 45366
diff changeset
   191
  end
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   192
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   193
fun minimize_facts do_learn prover_name (params as {learn, timeout, ...}) silent
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   194
                   i n state preplay0 facts =
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   195
  let
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40553
diff changeset
   196
    val ctxt = Proof.context_of state
45574
7a39df11bcf6 be more silent when auto minimizing
blanchet
parents: 45520
diff changeset
   197
    val prover =
7a39df11bcf6 be more silent when auto minimizing
blanchet
parents: 45520
diff changeset
   198
      get_prover ctxt (if silent then Auto_Minimize else Minimize) prover_name
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   199
    fun test timeout = test_facts params silent prover timeout i n state
48796
0f94b8b69e79 consider removing chained facts last, so that they're more likely to be kept
blanchet
parents: 48399
diff changeset
   200
    val (chained, non_chained) = List.partition is_fact_chained facts
48799
5c9356f8c968 warn users about unused "using" facts
blanchet
parents: 48798
diff changeset
   201
    (* Push chained facts to the back, so that they are less likely to be
5c9356f8c968 warn users about unused "using" facts
blanchet
parents: 48798
diff changeset
   202
       kicked out by the linear minimization algorithm. *)
48796
0f94b8b69e79 consider removing chained facts last, so that they're more likely to be kept
blanchet
parents: 48399
diff changeset
   203
    val facts = non_chained @ chained
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   204
  in
48796
0f94b8b69e79 consider removing chained facts last, so that they're more likely to be kept
blanchet
parents: 48399
diff changeset
   205
    (print silent (fn () => "Sledgehammer minimizer: " ^
0f94b8b69e79 consider removing chained facts last, so that they're more likely to be kept
blanchet
parents: 48399
diff changeset
   206
                            quote prover_name ^ ".");
0f94b8b69e79 consider removing chained facts last, so that they're more likely to be kept
blanchet
parents: 48399
diff changeset
   207
     case test timeout facts of
45371
blanchet
parents: 45370
diff changeset
   208
       result as {outcome = NONE, used_facts, run_time, ...} =>
38015
b30c3c2e1030 implemented "sublinear" minimization algorithm
blanchet
parents: 38002
diff changeset
   209
       let
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48796
diff changeset
   210
         val facts = filter_used_facts true used_facts facts
45372
cc455b2897f8 cascading timeouts in minimizer, part 2
blanchet
parents: 45371
diff changeset
   211
         val min =
42646
4781fcd53572 replaced some Unsynchronized.refs with Config.Ts
blanchet
parents: 42579
diff changeset
   212
           if length facts >= Config.get ctxt binary_min_facts then
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   213
             binary_minimize
40977
9140c5950494 [mq]: sledge_binary_minimizer
blanchet
parents: 40941
diff changeset
   214
           else
45368
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   215
             linear_minimize
ff2edf24e83a cascading timeouts in minimizer
blanchet
parents: 45367
diff changeset
   216
         val (min_facts, {preplay, message, message_tail, ...}) =
45371
blanchet
parents: 45370
diff changeset
   217
           min test (new_timeout timeout run_time) result facts
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   218
       in
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   219
         print silent (fn () => cat_lines
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   220
             ["Minimized to " ^ n_facts (map fst min_facts)] ^
48796
0f94b8b69e79 consider removing chained facts last, so that they're more likely to be kept
blanchet
parents: 48399
diff changeset
   221
              (case min_facts |> filter is_fact_chained |> length of
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   222
                 0 => ""
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   223
               | n => "\n(including " ^ string_of_int n ^ " chained)") ^ ".");
48399
4bacc8983b3d learn from SMT proofs when they can be minimized by Metis
blanchet
parents: 48396
diff changeset
   224
         (if learn then do_learn prover_name (maps snd min_facts) else ());
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   225
         (SOME min_facts,
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   226
          (if is_some preplay0 andalso length min_facts = length facts then
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   227
             the preplay0
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   228
           else
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   229
             preplay,
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   230
           message, message_tail))
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   231
       end
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   232
     | {outcome = SOME TimedOut, preplay, ...} =>
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   233
       (NONE,
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   234
        (preplay,
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   235
         fn _ =>
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   236
            "Timeout: You can increase the time limit using the \"timeout\" \
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   237
            \option (e.g., \"timeout = " ^
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   238
            string_of_int (10 + Time.toMilliseconds
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   239
                (timeout |> the_default (seconds 60.0)) div 1000) ^
31313171deb5 thread no timeout properly
blanchet
parents: 50020
diff changeset
   240
            "\").", ""))
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   241
     | {preplay, message, ...} =>
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
   242
       (NONE, (preplay, prefix "Prover error: " o message, "")))
43166
68e3cd19fee8 show what failed to play
blanchet
parents: 43085
diff changeset
   243
    handle ERROR msg =>
50669
84c7cf36b2e0 use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents: 50668
diff changeset
   244
           (NONE, (Lazy.value (Failed_to_Play plain_metis),
84c7cf36b2e0 use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents: 50668
diff changeset
   245
            fn _ => "Error: " ^ msg, ""))
31236
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   246
  end
2a1f5c87ac28 proper signature constraint;
wenzelm
parents: 31037
diff changeset
   247
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   248
fun adjust_reconstructor_params override_params
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
   249
        ({debug, verbose, overlord, spy, blocking, provers, type_enc, strict,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   250
         lam_trans, uncurried_aliases, learn, fact_filter, max_facts,
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49914
diff changeset
   251
         fact_thresholds, max_mono_iters, max_new_mono_instances, isar_proofs,
53764
eba0d1c069b8 merged "isar_try0" and "isar_minimize" options
blanchet
parents: 53763
diff changeset
   252
         isar_compress, isar_try0, slice, minimize, timeout, preplay_timeout,
eba0d1c069b8 merged "isar_try0" and "isar_minimize" options
blanchet
parents: 53763
diff changeset
   253
         expect} : params) =
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   254
  let
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   255
    fun lookup_override name default_value =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   256
      case AList.lookup (op =) override_params name of
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   257
        SOME [s] => SOME s
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   258
      | _ => default_value
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   259
    (* Only those options that reconstructors are interested in are considered
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   260
       here. *)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   261
    val type_enc = lookup_override "type_enc" type_enc
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   262
    val lam_trans = lookup_override "lam_trans" lam_trans
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   263
  in
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
   264
    {debug = debug, verbose = verbose, overlord = overlord, spy = spy, blocking = blocking,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
   265
     provers = provers, type_enc = type_enc, strict = strict, lam_trans = lam_trans,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
   266
     uncurried_aliases = uncurried_aliases, learn = learn, fact_filter = fact_filter,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
   267
     max_facts = max_facts, fact_thresholds = fact_thresholds, max_mono_iters = max_mono_iters,
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49914
diff changeset
   268
     max_new_mono_instances = max_new_mono_instances, isar_proofs = isar_proofs,
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
   269
     isar_compress = isar_compress, isar_try0 = isar_try0, slice = slice, minimize = minimize,
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
   270
     timeout = timeout, preplay_timeout = preplay_timeout, expect = expect}
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   271
  end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   272
48384
83dc102041e6 learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents: 48321
diff changeset
   273
fun maybe_minimize ctxt mode do_learn name
49918
cf441f4a358b renamed Isar-proof related options + changed semantics of Isar shrinking
blanchet
parents: 49914
diff changeset
   274
        (params as {verbose, isar_proofs, minimize, ...})
51010
afd0213a3dab tuned data structure
blanchet
parents: 51009
diff changeset
   275
        ({state, subgoal, subgoal_count, ...} : prover_problem)
51009
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51007
diff changeset
   276
        (result as {outcome, used_facts, used_from, run_time, preplay, message,
48384
83dc102041e6 learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents: 48321
diff changeset
   277
                    message_tail} : prover_result) =
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   278
  if is_some outcome orelse null used_facts then
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   279
    result
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   280
  else
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   281
    let
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   282
      val num_facts = length used_facts
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   283
      val ((perhaps_minimize, (minimize_name, params)), preplay) =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   284
        if mode = Normal then
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   285
          if num_facts >= Config.get ctxt auto_minimize_min_facts then
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   286
            ((true, (name, params)), preplay)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   287
          else
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   288
            let
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   289
              fun can_min_fast_enough time =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   290
                0.001
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   291
                * Real.fromInt ((num_facts + 1) * Time.toMilliseconds time)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   292
                <= Config.get ctxt auto_minimize_max_time
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   293
              fun prover_fast_enough () = can_min_fast_enough run_time
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   294
            in
51190
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   295
              (case Lazy.force preplay of
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   296
                 Played (reconstr as Metis _, timeout) =>
51191
89e9e945a005 auto-minimizer should respect "isar_proofs = true"
blanchet
parents: 51190
diff changeset
   297
                 if isar_proofs = SOME true then
51200
260cb10aac4b minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents: 51191
diff changeset
   298
                   (* Cheat: Assume the selected ATP is as fast as "metis" for
260cb10aac4b minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents: 51191
diff changeset
   299
                      the goal it proved itself. *)
260cb10aac4b minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents: 51191
diff changeset
   300
                   (can_min_fast_enough timeout,
260cb10aac4b minimize SMT proofs with E if Isar proofs are desired and Metis managed to preplay
blanchet
parents: 51191
diff changeset
   301
                    (isar_proof_reconstructor ctxt name, params))
51191
89e9e945a005 auto-minimizer should respect "isar_proofs = true"
blanchet
parents: 51190
diff changeset
   302
                 else if can_min_fast_enough timeout then
51190
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   303
                   (true, extract_reconstructor params reconstr
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   304
                          ||> (fn override_params =>
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   305
                                  adjust_reconstructor_params override_params
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   306
                                      params))
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   307
                 else
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   308
                   (prover_fast_enough (), (name, params))
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   309
               | Played (SMT, timeout) =>
51191
89e9e945a005 auto-minimizer should respect "isar_proofs = true"
blanchet
parents: 51190
diff changeset
   310
                 (* Cheat: Assume the original prover is as fast as "smt" for
51190
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   311
                    the goal it proved itself *)
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   312
                 (can_min_fast_enough timeout, (name, params))
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   313
               | _ => (prover_fast_enough (), (name, params)),
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51187
diff changeset
   314
               preplay)
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   315
            end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   316
        else
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   317
          ((false, (name, params)), preplay)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   318
      val minimize = minimize |> the_default perhaps_minimize
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   319
      val (used_facts, (preplay, message, _)) =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   320
        if minimize then
48392
ca998fa08cd9 added "learn_from_atp" command to MaSh, for patient users
blanchet
parents: 48384
diff changeset
   321
          minimize_facts do_learn minimize_name params
51005
ce4290c33d73 eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents: 50669
diff changeset
   322
              (mode <> Normal orelse not verbose) subgoal subgoal_count state
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   323
              (SOME preplay)
51010
afd0213a3dab tuned data structure
blanchet
parents: 51009
diff changeset
   324
              (filter_used_facts true used_facts (map (apsnd single) used_from))
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   325
          |>> Option.map (map fst)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   326
        else
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   327
          (SOME used_facts, (preplay, message, ""))
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   328
    in
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   329
      case used_facts of
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   330
        SOME used_facts =>
51009
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51007
diff changeset
   331
        {outcome = NONE, used_facts = used_facts, used_from = used_from,
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51007
diff changeset
   332
         run_time = run_time, preplay = preplay, message = message,
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51007
diff changeset
   333
         message_tail = message_tail}
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   334
      | NONE => result
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   335
    end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   336
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   337
fun get_minimizing_prover ctxt mode do_learn name params minimize_command
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   338
                          problem =
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   339
  get_prover ctxt mode name params minimize_command problem
48384
83dc102041e6 learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
blanchet
parents: 48321
diff changeset
   340
  |> maybe_minimize ctxt mode do_learn name params problem
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents: 48085
diff changeset
   341
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48314
diff changeset
   342
fun run_minimize (params as {provers, ...}) do_learn i refs state =
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   343
  let
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   344
    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
   345
    val reserved = reserved_isar_keyword_table ()
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48250
diff changeset
   346
    val chained_ths = #facts (Proof.goal state)
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   347
    val css = clasimpset_rule_table_of ctxt
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
   348
    val facts =
46388
e7445478d90b proper statuses for "fact_from_ref"
blanchet
parents: 46340
diff changeset
   349
      refs |> maps (map (apsnd single)
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   350
                    o fact_of_ref ctxt reserved chained_ths css)
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   351
  in
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   352
    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
   353
      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
   354
    | 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
   355
             [] => 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
   356
           | 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
   357
             (kill_provers ();
51187
c344cf148e8f avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
blanchet
parents: 51130
diff changeset
   358
              minimize_facts do_learn prover params false i n state NONE facts
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
   359
              |> (fn (_, (preplay, message, message_tail)) =>
50669
84c7cf36b2e0 use "Lazy" to simplify control flow a bit and guarantee single evaluation (at most)
blanchet
parents: 50668
diff changeset
   360
                     message (Lazy.force preplay) ^ message_tail
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43259
diff changeset
   361
                     |> Output.urgent_message))
38045
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   362
  end
f367847f5068 minor refactoring
blanchet
parents: 38023
diff changeset
   363
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35865
diff changeset
   364
end;