src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML
author blanchet
Fri, 01 Aug 2014 14:43:57 +0200
changeset 57732 e1b2442dc629
parent 57724 9ea51eddf81c
child 57734 18bb3e1ff6f6
permissions -rw-r--r--
simplified minimization logic
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55201
1ee776da8da7 renamed ML file
blanchet
parents: 55183
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_prover.ML
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
     2
    Author:     Fabian Immler, TU Muenchen
32996
d2e48879e65a removed disjunctive group cancellation -- provers run independently;
wenzelm
parents: 32995
diff changeset
     3
    Author:     Makarius
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
     5
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 41066
diff changeset
     6
Generic prover abstraction for Sledgehammer.
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
     7
*)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
     8
55201
1ee776da8da7 renamed ML file
blanchet
parents: 55183
diff changeset
     9
signature SLEDGEHAMMER_PROVER =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
    10
sig
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53551
diff changeset
    11
  type atp_failure = ATP_Proof.atp_failure
46340
cac402c486b0 separate orthogonal components
blanchet
parents: 46320
diff changeset
    12
  type stature = ATP_Problem_Generate.stature
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    13
  type type_enc = ATP_Problem_Generate.type_enc
51005
ce4290c33d73 eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents: 50927
diff changeset
    14
  type fact = Sledgehammer_Fact.fact
55287
ffa306239316 renamed ML file
blanchet
parents: 55286
diff changeset
    15
  type proof_method = Sledgehammer_Proof_Methods.proof_method
ffa306239316 renamed ML file
blanchet
parents: 55286
diff changeset
    16
  type play_outcome = Sledgehammer_Proof_Methods.play_outcome
ffa306239316 renamed ML file
blanchet
parents: 55286
diff changeset
    17
  type minimize_command = Sledgehammer_Proof_Methods.minimize_command
39493
cb2208f2c07d move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents: 39492
diff changeset
    18
53052
a0db255af8c5 sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
wenzelm
parents: 53047
diff changeset
    19
  datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize
43021
5910dd009d0e handle non-auto try case of Sledgehammer better
blanchet
parents: 43015
diff changeset
    20
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
    21
  type params =
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    22
    {debug : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    23
     verbose : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    24
     overlord : bool,
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
    25
     spy : bool,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    26
     blocking : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    27
     provers : string list,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    28
     type_enc : string option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    29
     strict : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    30
     lam_trans : string option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    31
     uncurried_aliases : bool option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    32
     learn : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    33
     fact_filter : string option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    34
     max_facts : int option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    35
     fact_thresholds : real * real,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    36
     max_mono_iters : int option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    37
     max_new_mono_instances : int option,
51190
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51186
diff changeset
    38
     isar_proofs : bool option,
57245
f6bf6d5341ee renamed Sledgehammer options
blanchet
parents: 57208
diff changeset
    39
     compress : real,
f6bf6d5341ee renamed Sledgehammer options
blanchet
parents: 57208
diff changeset
    40
     try0 : bool,
55297
1dfcd49f5dcb renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents: 55288
diff changeset
    41
     smt_proofs : bool option,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    42
     slice : bool,
57721
e4858f85e616 always minimize Sledgehammer results by default
blanchet
parents: 57718
diff changeset
    43
     minimize : bool,
54816
10d48c2a3e32 made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents: 54815
diff changeset
    44
     timeout : Time.time,
10d48c2a3e32 made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents: 54815
diff changeset
    45
     preplay_timeout : Time.time,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    46
     expect : string}
39493
cb2208f2c07d move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents: 39492
diff changeset
    47
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    48
  type prover_problem =
54141
f57f8e7a879f generate a comment storing the goal nickname in "learn_prover"
blanchet
parents: 54126
diff changeset
    49
    {comment : string,
f57f8e7a879f generate a comment storing the goal nickname in "learn_prover"
blanchet
parents: 54126
diff changeset
    50
     state : Proof.state,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    51
     goal : thm,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    52
     subgoal : int,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    53
     subgoal_count : int,
51010
afd0213a3dab tuned data structure
blanchet
parents: 51009
diff changeset
    54
     factss : (string * fact list) list}
39493
cb2208f2c07d move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents: 39492
diff changeset
    55
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
    56
  type prover_result =
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53551
diff changeset
    57
    {outcome : atp_failure option,
51009
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51007
diff changeset
    58
     used_facts : (string * stature) list,
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51007
diff changeset
    59
     used_from : fact list,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    60
     run_time : Time.time,
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
    61
     preplay : (proof_method * play_outcome) Lazy.lazy,
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
    62
     message : proof_method * play_outcome -> string,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
    63
     message_tail : string}
39493
cb2208f2c07d move SPASS's Flotter hack to "Sledgehammer_Reconstruct"
blanchet
parents: 39492
diff changeset
    64
43051
d7075adac3bd minimize with Metis if possible
blanchet
parents: 43050
diff changeset
    65
  type prover =
57056
8b2283566f6e properly reconstruct helpers in Z3 proofs
blanchet
parents: 57054
diff changeset
    66
    params -> ((string * string list) list -> string -> minimize_command) -> prover_problem ->
8b2283566f6e properly reconstruct helpers in Z3 proofs
blanchet
parents: 57054
diff changeset
    67
    prover_result
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
    68
48319
340187063d84 use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents: 48314
diff changeset
    69
  val SledgehammerN : string
57037
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
    70
  val str_of_mode : mode -> string
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
    71
  val overlord_file_location_of_prover : string -> string * string
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
    72
  val proof_banner : mode -> string -> string
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43044
diff changeset
    73
  val is_atp : theory -> string -> bool
55288
1a4358d14ce2 added 'smt' option to control generation of 'by smt' proofs
blanchet
parents: 55287
diff changeset
    74
  val bunch_of_proof_methods : bool -> bool -> string -> proof_method list
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48797
diff changeset
    75
  val is_fact_chained : (('a * stature) * 'b) -> bool
57056
8b2283566f6e properly reconstruct helpers in Z3 proofs
blanchet
parents: 57054
diff changeset
    76
  val filter_used_facts : bool -> (''a * stature) list -> ((''a * stature) * 'b) list ->
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48797
diff changeset
    77
    ((''a * stature) * 'b) list
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
    78
  val play_one_line_proof : mode -> Time.time -> ((string * 'a) * thm) list -> Proof.state -> int ->
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
    79
    proof_method -> proof_method list -> proof_method * play_outcome
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
    80
  val repair_monomorph_context : int option -> int -> int option -> int -> Proof.context ->
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
    81
    Proof.context
55212
blanchet
parents: 55211
diff changeset
    82
blanchet
parents: 55211
diff changeset
    83
  val supported_provers : Proof.context -> unit
blanchet
parents: 55211
diff changeset
    84
  val kill_provers : unit -> unit
blanchet
parents: 55211
diff changeset
    85
  val running_provers : unit -> unit
blanchet
parents: 55211
diff changeset
    86
  val messages : int option -> unit
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
    87
end;
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
    88
55201
1ee776da8da7 renamed ML file
blanchet
parents: 55183
diff changeset
    89
structure Sledgehammer_Prover : SLEDGEHAMMER_PROVER =
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
    90
struct
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff changeset
    91
57154
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 57056
diff changeset
    92
open ATP_Proof
43085
0a2f5b86bdd7 first step in sharing more code between ATP and Metis translation
blanchet
parents: 43063
diff changeset
    93
open ATP_Util
57154
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 57056
diff changeset
    94
open ATP_Systems
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38023
diff changeset
    95
open ATP_Problem
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    96
open ATP_Problem_Generate
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    97
open ATP_Proof_Reconstruct
45521
0cd6e59bd0b5 give each time slice its own lambda translation
blanchet
parents: 45520
diff changeset
    98
open Metis_Tactic
51005
ce4290c33d73 eliminated needless speed optimization -- and simplified code quite a bit
blanchet
parents: 50927
diff changeset
    99
open Sledgehammer_Fact
55287
ffa306239316 renamed ML file
blanchet
parents: 55286
diff changeset
   100
open Sledgehammer_Proof_Methods
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   101
open Sledgehammer_Isar_Proof
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   102
open Sledgehammer_Isar_Preplay
54000
9cfff7f61d0d added experimental configuration options to tune use of builtin symbols in SMT
blanchet
parents: 53989
diff changeset
   103
45376
blanchet
parents: 45370
diff changeset
   104
(* Identifier that distinguishes Sledgehammer from other tools that could use
38102
019a49759829 fix bug in the newly introduced "bound concealing" code
blanchet
parents: 38100
diff changeset
   105
   "Async_Manager". *)
48319
340187063d84 use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents: 48314
diff changeset
   106
val SledgehammerN = "Sledgehammer"
37585
c2ed8112ce57 multiplexing
blanchet
parents: 37584
diff changeset
   107
57037
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   108
datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   109
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   110
fun str_of_mode Auto_Try = "Auto Try"
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   111
  | str_of_mode Try = "Try"
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   112
  | str_of_mode Normal = "Normal"
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   113
  | str_of_mode MaSh = "MaSh"
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   114
  | str_of_mode Auto_Minimize = "Auto_Minimize"
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   115
  | str_of_mode Minimize = "Minimize"
c51132be8e16 avoid markup-generating @{make_string}
blanchet
parents: 56985
diff changeset
   116
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43044
diff changeset
   117
val is_atp = member (op =) o supported_atps
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43044
diff changeset
   118
35969
c9565298df9e added support for Sledgehammer parameters;
blanchet
parents: 35867
diff changeset
   119
type params =
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   120
  {debug : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   121
   verbose : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   122
   overlord : bool,
53800
ac1ec5065316 added "spy" option to Sledgehammer
blanchet
parents: 53764
diff changeset
   123
   spy : bool,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   124
   blocking : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   125
   provers : string list,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   126
   type_enc : string option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   127
   strict : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   128
   lam_trans : string option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   129
   uncurried_aliases : bool option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   130
   learn : bool,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   131
   fact_filter : string option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   132
   max_facts : int option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   133
   fact_thresholds : real * real,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   134
   max_mono_iters : int option,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   135
   max_new_mono_instances : int option,
51190
2654b3965c8d made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
blanchet
parents: 51186
diff changeset
   136
   isar_proofs : bool option,
57245
f6bf6d5341ee renamed Sledgehammer options
blanchet
parents: 57208
diff changeset
   137
   compress : real,
f6bf6d5341ee renamed Sledgehammer options
blanchet
parents: 57208
diff changeset
   138
   try0 : bool,
55297
1dfcd49f5dcb renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents: 55288
diff changeset
   139
   smt_proofs : bool option,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   140
   slice : bool,
57721
e4858f85e616 always minimize Sledgehammer results by default
blanchet
parents: 57718
diff changeset
   141
   minimize : bool,
54816
10d48c2a3e32 made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents: 54815
diff changeset
   142
   timeout : Time.time,
10d48c2a3e32 made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents: 54815
diff changeset
   143
   preplay_timeout : Time.time,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   144
   expect : string}
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
   145
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   146
type prover_problem =
54141
f57f8e7a879f generate a comment storing the goal nickname in "learn_prover"
blanchet
parents: 54126
diff changeset
   147
  {comment : string,
f57f8e7a879f generate a comment storing the goal nickname in "learn_prover"
blanchet
parents: 54126
diff changeset
   148
   state : Proof.state,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   149
   goal : thm,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   150
   subgoal : int,
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   151
   subgoal_count : int,
51010
afd0213a3dab tuned data structure
blanchet
parents: 51009
diff changeset
   152
   factss : (string * fact list) list}
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
   153
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   154
type prover_result =
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 53551
diff changeset
   155
  {outcome : atp_failure option,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   156
   used_facts : (string * stature) list,
51009
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51007
diff changeset
   157
   used_from : fact list,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   158
   run_time : Time.time,
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   159
   preplay : (proof_method * play_outcome) Lazy.lazy,
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   160
   message : proof_method * play_outcome -> string,
48321
c552d7f1720b learn from minimized ATP proofs
blanchet
parents: 48319
diff changeset
   161
   message_tail : string}
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
   162
43051
d7075adac3bd minimize with Metis if possible
blanchet
parents: 43050
diff changeset
   163
type prover =
45520
2b1dde0b1c30 thread in additional options to minimizer
blanchet
parents: 45519
diff changeset
   164
  params -> ((string * string list) list -> string -> minimize_command)
2b1dde0b1c30 thread in additional options to minimizer
blanchet
parents: 45519
diff changeset
   165
  -> prover_problem -> prover_result
35867
16279c4c7a33 move all ATP setup code into ATP_Wrapper
blanchet
parents: 35866
diff changeset
   166
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   167
fun overlord_file_location_of_prover prover = (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
41313
a96ac4d180b7 optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents: 41259
diff changeset
   168
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   169
fun proof_banner mode name =
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   170
  (case mode of
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43031
diff changeset
   171
    Auto_Try => "Auto Sledgehammer (" ^ quote name ^ ") found a proof"
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43031
diff changeset
   172
  | Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   173
  | _ => "Try this")
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43031
diff changeset
   174
55297
1dfcd49f5dcb renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
blanchet
parents: 55288
diff changeset
   175
fun bunch_of_proof_methods smt_proofs needs_full_types desperate_lam_trans =
57724
9ea51eddf81c put faster proof methods first
blanchet
parents: 57723
diff changeset
   176
  [Simp_Method, Auto_Method, Fastforce_Method, Blast_Method, Meson_Method,
9ea51eddf81c put faster proof methods first
blanchet
parents: 57723
diff changeset
   177
   Metis_Method (if needs_full_types then SOME full_typesN else NONE, NONE),
9ea51eddf81c put faster proof methods first
blanchet
parents: 57723
diff changeset
   178
   Force_Method, Linarith_Method, Presburger_Method] @
55288
1a4358d14ce2 added 'smt' option to control generation of 'by smt' proofs
blanchet
parents: 55287
diff changeset
   179
  (if needs_full_types then
57718
892e8e7a42b3 added more proof methods for one-liners
blanchet
parents: 57713
diff changeset
   180
     [Metis_Method (SOME really_full_type_enc, NONE),
55345
8a53ee72e595 try right bunch of methods
blanchet
parents: 55323
diff changeset
   181
      Metis_Method (SOME full_typesN, SOME desperate_lam_trans),
8a53ee72e595 try right bunch of methods
blanchet
parents: 55323
diff changeset
   182
      Metis_Method (SOME really_full_type_enc, SOME desperate_lam_trans)]
55288
1a4358d14ce2 added 'smt' option to control generation of 'by smt' proofs
blanchet
parents: 55287
diff changeset
   183
   else
57718
892e8e7a42b3 added more proof methods for one-liners
blanchet
parents: 57713
diff changeset
   184
     [Metis_Method (SOME full_typesN, NONE),
55345
8a53ee72e595 try right bunch of methods
blanchet
parents: 55323
diff changeset
   185
      Metis_Method (SOME no_typesN, SOME desperate_lam_trans),
8a53ee72e595 try right bunch of methods
blanchet
parents: 55323
diff changeset
   186
      Metis_Method (SOME really_full_type_enc, SOME desperate_lam_trans)]) @
56081
72fad75baf7e integrate SMT2 with Sledgehammer
blanchet
parents: 55475
diff changeset
   187
  (if smt_proofs then [SMT2_Method] else [])
45561
57227eedce81 don't propagate user-set "type_enc" or "lam_trans" to Metis calls
blanchet
parents: 45560
diff changeset
   188
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   189
fun preplay_methods timeout facts meths state i =
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43031
diff changeset
   190
  let
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   191
    val {context = ctxt, facts = chained, goal} = Proof.goal state
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   192
    val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   193
    val step =
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   194
      Prove ([], [], ("", 0), Logic.list_implies (map prop_of chained @ hyp_ts, concl_t), [],
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   195
        ([], facts), meths, "")
54816
10d48c2a3e32 made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents: 54815
diff changeset
   196
  in
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   197
    preplay_isar_step ctxt timeout [] step
54816
10d48c2a3e32 made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents: 54815
diff changeset
   198
  end
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43031
diff changeset
   199
48798
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48797
diff changeset
   200
fun is_fact_chained ((_, (sc, _)), _) = sc = Chained
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48797
diff changeset
   201
9152e66f98da be less aggressive at kicking out chained facts
blanchet
parents: 48797
diff changeset
   202
fun filter_used_facts keep_chained used =
54773
79f66cd15d57 fixed source of 'Subscript' exception
blanchet
parents: 54772
diff changeset
   203
  filter ((member (op =) used o fst) orf (if keep_chained then is_fact_chained else K false))
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43031
diff changeset
   204
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   205
fun play_one_line_proof mode timeout used_facts state i preferred (meths as meth :: _) =
43034
18259246abb5 try both "metis" and (on failure) "metisFT" in replay
blanchet
parents: 43033
diff changeset
   206
  let
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   207
    fun dont_know () =
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   208
      (if member (op =) meths preferred then preferred else meth, Play_Timed_Out Time.zeroTime)
43034
18259246abb5 try both "metis" and (on failure) "metisFT" in replay
blanchet
parents: 43033
diff changeset
   209
  in
54816
10d48c2a3e32 made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents: 54815
diff changeset
   210
    if timeout = Time.zeroTime then
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   211
      dont_know ()
45379
0147a4348ca1 try "smt" as a fallback for ATPs if "metis" fails/times out
blanchet
parents: 45378
diff changeset
   212
    else
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50494
diff changeset
   213
      let
54815
4f6ec8754bf5 simplified data structure
blanchet
parents: 54813
diff changeset
   214
        val _ = if mode = Minimize then Output.urgent_message "Preplaying proof..." else ()
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   215
        val gfs = used_facts |> map (fst o fst) |> sort string_ord
54824
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
   216
      in
57723
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   217
        (case preplay_methods timeout gfs meths state i of
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   218
          res :: _ => res
668322cd58f4 use parallel preplay machinery also for one-line proofs
blanchet
parents: 57721
diff changeset
   219
        | [] => dont_know ())
54824
4e58a38b330b refactored preplaying outcome data structure
blanchet
parents: 54823
diff changeset
   220
      end
43033
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43031
diff changeset
   221
  end
c4b9b4be90c4 show time taken for reconstruction
blanchet
parents: 43031
diff changeset
   222
53480
247817dbb990 limit the number of instances of a single theorem
blanchet
parents: 53478
diff changeset
   223
val max_fact_instances = 10 (* FUDGE *)
247817dbb990 limit the number of instances of a single theorem
blanchet
parents: 53478
diff changeset
   224
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   225
fun repair_monomorph_context max_iters best_max_iters max_new_instances best_max_new_instances =
52034
11b48e7a4e7e correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents: 52031
diff changeset
   226
  Config.put Monomorph.max_rounds (max_iters |> the_default best_max_iters)
11b48e7a4e7e correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents: 52031
diff changeset
   227
  #> Config.put Monomorph.max_new_instances
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   228
       (max_new_instances |> the_default best_max_new_instances)
53480
247817dbb990 limit the number of instances of a single theorem
blanchet
parents: 53478
diff changeset
   229
  #> Config.put Monomorph.max_thm_instances max_fact_instances
52034
11b48e7a4e7e correctly 'repair' the monomorphization context for SMT solvers from Sledgehammer
blanchet
parents: 52031
diff changeset
   230
55212
blanchet
parents: 55211
diff changeset
   231
fun supported_provers ctxt =
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43044
diff changeset
   232
  let
55212
blanchet
parents: 55211
diff changeset
   233
    val thy = Proof_Context.theory_of ctxt
blanchet
parents: 55211
diff changeset
   234
    val (remote_provers, local_provers) =
blanchet
parents: 55211
diff changeset
   235
      sort_strings (supported_atps thy) @
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56093
diff changeset
   236
      sort_strings (SMT2_Config.available_solvers_of ctxt)
55212
blanchet
parents: 55211
diff changeset
   237
      |> List.partition (String.isPrefix remote_prefix)
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43044
diff changeset
   238
  in
57056
8b2283566f6e properly reconstruct helpers in Z3 proofs
blanchet
parents: 57054
diff changeset
   239
    Output.urgent_message ("Supported provers: " ^ commas (local_provers @ remote_provers) ^ ".")
43050
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43044
diff changeset
   240
  end
59284a13abc4 support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents: 43044
diff changeset
   241
55212
blanchet
parents: 55211
diff changeset
   242
fun kill_provers () = Async_Manager.kill_threads SledgehammerN "prover"
blanchet
parents: 55211
diff changeset
   243
fun running_provers () = Async_Manager.running_threads SledgehammerN "prover"
blanchet
parents: 55211
diff changeset
   244
val messages = Async_Manager.thread_messages SledgehammerN "prover"
blanchet
parents: 55211
diff changeset
   245
28582
c269a3045fdf info: back to plain printing;
wenzelm
parents: 28571
diff changeset
   246
end;