src/HOL/Tools/SMT/smt_config.ML
author wenzelm
Fri, 05 Mar 2021 20:38:55 +0100
changeset 73388 a40e69fde2b4
parent 72908 6a26a955308e
child 73389 f3378101f555
permissions -rw-r--r--
clarified smt: support Timeout.ignored and Timeout.scale_time;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
     1
(*  Title:      HOL/Tools/SMT/smt_config.ML
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     3
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     4
Configuration options and diagnostic tools for SMT.
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     5
*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     6
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
     7
signature SMT_CONFIG =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     8
sig
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     9
  (*solver*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    10
  type solver_info = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    11
    name: string,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
    12
    class: Proof.context -> SMT_Util.class,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    13
    avail: unit -> bool,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    14
    options: Proof.context -> string list }
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    15
  val add_solver: solver_info -> Context.generic -> Context.generic
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    16
  val set_solver_options: string * string -> Context.generic -> Context.generic
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
  val is_available: Proof.context -> string -> bool
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    18
  val available_solvers_of: Proof.context -> string list
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    19
  val select_solver: string -> Context.generic -> Context.generic
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    20
  val solver_of: Proof.context -> string
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
    21
  val solver_class_of: Proof.context -> SMT_Util.class
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    22
  val solver_options_of: Proof.context -> string list
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    23
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    24
  (*options*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    25
  val oracle: bool Config.T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    26
  val timeout: real Config.T
59213
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
    27
  val reconstruction_step_timeout: real Config.T
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    28
  val random_seed: int Config.T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    29
  val read_only_certificates: bool Config.T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    30
  val verbose: bool Config.T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
  val trace: bool Config.T
59215
35c13f4995b1 optionally display statistics for Z3 proof reconstruction
boehmes
parents: 59214
diff changeset
    32
  val statistics: bool Config.T
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    33
  val monomorph_limit: int Config.T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    34
  val monomorph_instances: int Config.T
66738
793e7a9c30c5 properly take quantifiers into account (cf. my Ph.D. thesis, Section 6.4.1) and offer three modes of completeness (for experiments mostly)
blanchet
parents: 66551
diff changeset
    35
  val explicit_application: int Config.T
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66298
diff changeset
    36
  val higher_order: bool Config.T
66298
5ff9fe3fee66 introduced option for nat-as-int in SMT
blanchet
parents: 62519
diff changeset
    37
  val nat_as_int: bool Config.T
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    38
  val infer_triggers: bool Config.T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    39
  val debug_files: string Config.T
56816
2f3756ccba41 use internal proof-producing SAT solver for more efficient SMT proof replay
boehmes
parents: 56294
diff changeset
    40
  val sat_solver: string Config.T
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
    41
  val compress_verit_proofs: Proof.context -> bool
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    42
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    43
  (*tools*)
73388
a40e69fde2b4 clarified smt: support Timeout.ignored and Timeout.scale_time;
wenzelm
parents: 72908
diff changeset
    44
  val get_timeout: Proof.context -> Time.time option
59213
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
    45
  val with_time_limit: Proof.context -> real Config.T -> ('a -> 'b) -> 'a -> 'b
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    46
  val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    47
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    48
  (*diagnostics*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    49
  val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    50
  val verbose_msg: Proof.context -> ('a -> string) -> 'a -> unit
59215
35c13f4995b1 optionally display statistics for Z3 proof reconstruction
boehmes
parents: 59214
diff changeset
    51
  val statistics_msg: Proof.context -> ('a -> string) -> 'a -> unit
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
    52
  val verit_msg: Proof.context -> (unit -> 'a) -> unit
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
    53
  val verit_arith_msg: Proof.context -> (unit -> 'a) -> unit
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
    54
  val spy_verit: Proof.context -> bool
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
    55
  val spy_Z3: Proof.context -> bool
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    56
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
  (*certificates*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    58
  val select_certificates: string -> Context.generic -> Context.generic
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
  val certificates_of: Proof.context -> Cache_IO.cache option
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    60
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    61
  (*setup*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    62
  val print_setup: Proof.context -> unit
57229
blanchet
parents: 57209
diff changeset
    63
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    64
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
    65
structure SMT_Config: SMT_CONFIG =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    67
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    68
(* solver *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    69
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    70
type solver_info = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    71
  name: string,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
    72
  class: Proof.context -> SMT_Util.class,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    73
  avail: unit -> bool,
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56078
diff changeset
    74
  options: Proof.context -> string list}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    75
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    76
type data = {
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    77
  solvers: (solver_info * string list) Symtab.table,
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    78
  solver: string option,
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    79
  certs: Cache_IO.cache option}
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    80
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    81
fun mk_data solvers solver certs: data = {solvers=solvers, solver=solver, certs=certs}
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    82
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    83
val empty_data = mk_data Symtab.empty NONE NONE
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    84
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    85
fun solvers_of ({solvers, ...}: data) = solvers
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    86
fun solver_of ({solver, ...}: data) = solver
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    87
fun certs_of ({certs, ...}: data) = certs
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    88
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    89
fun map_solvers f ({solvers, solver, certs}: data) = mk_data (f solvers) solver certs
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    90
fun map_solver f ({solvers, solver, certs}: data) = mk_data solvers (f solver) certs
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    91
fun put_certs c ({solvers, solver, ...}: data) = mk_data solvers solver c
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    92
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    93
fun merge_data ({solvers=ss1,solver=s1,certs=c1}: data, {solvers=ss2,solver=s2,certs=c2}: data) =
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    94
  mk_data (Symtab.merge (K true) (ss1, ss2)) (merge_options (s1, s2)) (merge_options (c1, c2))
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    95
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    96
structure Data = Generic_Data
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    97
(
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    98
  type T = data
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    99
  val empty = empty_data
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   100
  val extend = I
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   101
  val merge = merge_data
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   102
)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   103
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   104
fun set_solver_options (name, options) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
  let val opts = String.tokens (Symbol.is_ascii_blank o str) options
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   106
  in Data.map (map_solvers (Symtab.map_entry name (apsnd (K opts)))) end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   108
fun add_solver (info as {name, ...} : solver_info) context =
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   109
  if Symtab.defined (solvers_of (Data.get context)) name then
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   110
    error ("Solver already registered: " ^ quote name)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   111
  else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   112
    context
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   113
    |> Data.map (map_solvers (Symtab.update (name, (info, []))))
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   114
    |> Context.map_theory (Attrib.setup (Binding.name (name ^ "_options"))
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   115
        (Scan.lift (\<^keyword>\<open>=\<close> |-- Args.name) >>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   116
          (Thm.declaration_attribute o K o set_solver_options o pair name))
60312
ee6f9a97205d tuned message;
wenzelm
parents: 59936
diff changeset
   117
        ("additional command line options for SMT solver " ^ quote name))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   118
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   119
fun all_solvers_of ctxt = Symtab.keys (solvers_of (Data.get (Context.Proof ctxt)))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   120
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   121
fun solver_name_of ctxt = solver_of (Data.get (Context.Proof ctxt))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   122
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   123
fun is_available ctxt name =
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   124
  (case Symtab.lookup (solvers_of (Data.get (Context.Proof ctxt))) name of
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   125
    SOME ({avail, ...}, _) => avail ()
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   126
  | NONE => false)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   127
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   128
fun available_solvers_of ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   129
  filter (is_available ctxt) (all_solvers_of ctxt)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   130
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   131
fun warn_solver (Context.Proof ctxt) name =
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56208
diff changeset
   132
      if Context_Position.is_visible ctxt then
57199
blanchet
parents: 57157
diff changeset
   133
        warning ("The SMT solver " ^ quote name ^ " is not installed")
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56208
diff changeset
   134
      else ()
57199
blanchet
parents: 57157
diff changeset
   135
  | warn_solver _ _ = ()
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   136
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   137
fun select_solver name context =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   138
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   139
    val ctxt = Context.proof_of context
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   140
    val upd = Data.map (map_solver (K (SOME name)))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   141
  in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   142
    if not (member (op =) (all_solvers_of ctxt) name) then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   143
      error ("Trying to select unknown solver: " ^ quote name)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   144
    else if not (is_available ctxt name) then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   145
      (warn_solver context name; upd context)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   146
    else upd context
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   147
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   148
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   149
fun no_solver_err () = error "No SMT solver selected"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   150
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   151
fun solver_of ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   152
  (case solver_name_of ctxt of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   153
    SOME name => name
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   154
  | NONE => no_solver_err ())
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   155
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   156
fun solver_info_of default select ctxt =
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   157
  (case solver_name_of ctxt of
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   158
    NONE => default ()
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   159
  | SOME name => select (Symtab.lookup (solvers_of (Data.get (Context.Proof ctxt))) name))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   160
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   161
fun solver_class_of ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   162
  let fun class_of ({class, ...}: solver_info, _) = class ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   163
  in solver_info_of no_solver_err (class_of o the) ctxt end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   164
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
fun solver_options_of ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   167
    fun all_options NONE = []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   168
      | all_options (SOME ({options, ...} : solver_info, opts)) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   169
          opts @ options ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   170
  in solver_info_of (K []) all_options ctxt end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   171
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   172
val setup_solver =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   173
  Attrib.setup \<^binding>\<open>smt_solver\<close>
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   174
    (Scan.lift (\<^keyword>\<open>=\<close> |-- Args.name) >>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   175
      (Thm.declaration_attribute o K o select_solver))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   176
    "SMT solver configuration"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   177
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   178
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   179
(* options *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   180
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   181
val oracle = Attrib.setup_config_bool \<^binding>\<open>smt_oracle\<close> (K true)
72343
478b7599a1a0 Effectively disable timeout for smt method/tactic
desharna
parents: 69205
diff changeset
   182
val timeout = Attrib.setup_config_real \<^binding>\<open>smt_timeout\<close> (K 1000000.0)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   183
val reconstruction_step_timeout = Attrib.setup_config_real \<^binding>\<open>smt_reconstruction_step_timeout\<close> (K 10.0)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   184
val random_seed = Attrib.setup_config_int \<^binding>\<open>smt_random_seed\<close> (K 1)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   185
val read_only_certificates = Attrib.setup_config_bool \<^binding>\<open>smt_read_only_certificates\<close> (K false)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   186
val verbose = Attrib.setup_config_bool \<^binding>\<open>smt_verbose\<close> (K true)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   187
val trace = Attrib.setup_config_bool \<^binding>\<open>smt_trace\<close> (K false)
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   188
val trace_verit = Attrib.setup_config_bool \<^binding>\<open>smt_debug_verit\<close> (K false)
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   189
val spy_verit_attr = Attrib.setup_config_bool \<^binding>\<open>smt_spy_verit\<close> (K false)
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   190
val spy_z3 = Attrib.setup_config_bool \<^binding>\<open>smt_spy_z3\<close> (K false)
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   191
val trace_arith_verit = Attrib.setup_config_bool \<^binding>\<open>smt_debug_arith_verit\<close> (K false)
72908
6a26a955308e improve and activate compression for veriT proof reconstruction
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72511
diff changeset
   192
val trace_verit_compression = Attrib.setup_config_bool \<^binding>\<open>verit_compress_proofs\<close> (K true)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   193
val statistics = Attrib.setup_config_bool \<^binding>\<open>smt_statistics\<close> (K false)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   194
val monomorph_limit = Attrib.setup_config_int \<^binding>\<open>smt_monomorph_limit\<close> (K 10)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   195
val monomorph_instances = Attrib.setup_config_int \<^binding>\<open>smt_monomorph_instances\<close> (K 500)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   196
val explicit_application = Attrib.setup_config_int \<^binding>\<open>smt_explicit_application\<close> (K 1)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   197
val higher_order = Attrib.setup_config_bool \<^binding>\<open>smt_higher_order\<close> (K false)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   198
val nat_as_int = Attrib.setup_config_bool \<^binding>\<open>smt_nat_as_int\<close> (K false)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   199
val infer_triggers = Attrib.setup_config_bool \<^binding>\<open>smt_infer_triggers\<close> (K false)
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   200
val debug_files = Attrib.setup_config_string \<^binding>\<open>smt_debug_files\<close> (K "")
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   201
val sat_solver = Attrib.setup_config_string \<^binding>\<open>smt_sat_solver\<close> (K "cdclite")
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   202
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   203
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   204
(* diagnostics *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   205
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   206
fun cond_trace flag f x = if flag then tracing ("SMT: " ^ f x) else ()
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   207
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   208
fun verbose_msg ctxt = cond_trace (Config.get ctxt verbose)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   209
fun trace_msg ctxt = cond_trace (Config.get ctxt trace)
59215
35c13f4995b1 optionally display statistics for Z3 proof reconstruction
boehmes
parents: 59214
diff changeset
   210
fun statistics_msg ctxt = cond_trace (Config.get ctxt statistics)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   211
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   212
fun verit_msg ctxt (x : unit -> 'a) = if (Config.get ctxt trace_verit) then ignore(x ()) else ()
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   213
fun verit_arith_msg ctxt (x : unit -> 'a) = if (Config.get ctxt trace_arith_verit) then ignore(x ()) else ()
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   214
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   215
fun spy_verit ctxt  = Config.get ctxt spy_verit_attr
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   216
fun spy_Z3 ctxt  = Config.get ctxt spy_z3
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72343
diff changeset
   217
fun compress_verit_proofs ctxt  = Config.get ctxt trace_verit_compression
69205
8050734eee3e add reconstruction by veriT in method smt
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 67149
diff changeset
   218
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   219
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   220
(* tools *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   221
73388
a40e69fde2b4 clarified smt: support Timeout.ignored and Timeout.scale_time;
wenzelm
parents: 72908
diff changeset
   222
fun get_timeout ctxt =
a40e69fde2b4 clarified smt: support Timeout.ignored and Timeout.scale_time;
wenzelm
parents: 72908
diff changeset
   223
  let val t = seconds (Config.get ctxt timeout);
a40e69fde2b4 clarified smt: support Timeout.ignored and Timeout.scale_time;
wenzelm
parents: 72908
diff changeset
   224
  in if Timeout.ignored t then NONE else SOME (Timeout.scale_time t) end;
a40e69fde2b4 clarified smt: support Timeout.ignored and Timeout.scale_time;
wenzelm
parents: 72908
diff changeset
   225
59213
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
   226
fun with_time_limit ctxt timeout_config f x =
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 60312
diff changeset
   227
  Timeout.apply (seconds (Config.get ctxt timeout_config)) f x
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 60312
diff changeset
   228
    handle Timeout.TIMEOUT _ => raise SMT_Failure.SMT SMT_Failure.Time_Out
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   229
59213
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
   230
fun with_timeout ctxt = with_time_limit ctxt timeout
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
   231
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   232
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   233
(* certificates *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   234
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   235
val certificates_of = certs_of o Data.get o Context.Proof
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   236
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   237
val get_certificates_path = Option.map (Cache_IO.cache_path_of) o certificates_of
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   238
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   239
fun select_certificates name context = context |> Data.map (put_certs (
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   240
  if name = "" then NONE
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   241
  else
72511
460d743010bc clarified signature: overloaded "+" for Path.append;
wenzelm
parents: 72458
diff changeset
   242
    (Resources.master_directory (Context.theory_of context) + Path.explode name)
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   243
    |> SOME o Cache_IO.unsynchronized_init))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   244
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   245
val setup_certificates =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   246
  Attrib.setup \<^binding>\<open>smt_certificates\<close>
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   247
    (Scan.lift (\<^keyword>\<open>=\<close> |-- Args.name) >>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   248
      (Thm.declaration_attribute o K o select_certificates))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   249
    "SMT certificates configuration"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   250
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   251
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   252
(* setup *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   253
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   254
val _ = Theory.setup (
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   255
  setup_solver #>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   256
  setup_certificates)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   257
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   258
fun print_setup ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   259
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   260
    fun string_of_bool b = if b then "true" else "false"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   261
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   262
    val names = available_solvers_of ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   263
    val ns = if null names then ["(none)"] else sort_strings names
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   264
    val n = the_default "(none)" (solver_name_of ctxt)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   265
    val opts = solver_options_of ctxt
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   266
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   267
    val t = string_of_real (Config.get ctxt timeout)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   268
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   269
    val certs_filename =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   270
      (case get_certificates_path ctxt of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   271
        SOME path => Path.print path
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   272
      | NONE => "(disabled)")
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   273
  in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   274
    Pretty.writeln (Pretty.big_list "SMT setup:" [
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   275
      Pretty.str ("Current SMT solver: " ^ n),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   276
      Pretty.str ("Current SMT solver options: " ^ space_implode " " opts),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   277
      Pretty.str_list "Available SMT solvers: "  "" ns,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   278
      Pretty.str ("Current timeout: " ^ t ^ " seconds"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   279
      Pretty.str ("With proofs: " ^
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   280
        string_of_bool (not (Config.get ctxt oracle))),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   281
      Pretty.str ("Certificates cache: " ^ certs_filename),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   282
      Pretty.str ("Fixed certificates: " ^
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   283
        string_of_bool (Config.get ctxt read_only_certificates))])
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   284
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   285
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   286
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 66738
diff changeset
   287
  Outer_Syntax.command \<^command_keyword>\<open>smt_status\<close>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   288
    "show the available SMT solvers, the currently selected SMT solver, \
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   289
    \and the values of SMT configuration options"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   290
    (Scan.succeed (Toplevel.keep (print_setup o Toplevel.context_of)))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   291
57229
blanchet
parents: 57209
diff changeset
   292
end;