src/HOL/Tools/SMT/smt_config.ML
author wenzelm
Mon, 06 Apr 2015 17:06:48 +0200
changeset 59936 b8ffc3dc9e24
parent 59215 35c13f4995b1
child 60312 ee6f9a97205d
permissions -rw-r--r--
@{command_spec} is superseded by @{command_keyword};
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
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
  val infer_triggers: bool Config.T
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    36
  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
    37
  val sat_solver: string Config.T
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    38
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    39
  (*tools*)
59213
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
    40
  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
    41
  val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b
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
  (*diagnostics*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    44
  val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    45
  val verbose_msg: Proof.context -> ('a -> string) -> 'a -> unit
59215
35c13f4995b1 optionally display statistics for Z3 proof reconstruction
boehmes
parents: 59214
diff changeset
    46
  val statistics_msg: Proof.context -> ('a -> string) -> 'a -> unit
56078
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
  (*certificates*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    49
  val select_certificates: string -> Context.generic -> Context.generic
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    50
  val certificates_of: Proof.context -> Cache_IO.cache option
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    51
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
  (*setup*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    53
  val print_setup: Proof.context -> unit
57229
blanchet
parents: 57209
diff changeset
    54
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    55
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
    56
structure SMT_Config: SMT_CONFIG =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    58
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
(* solver *)
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
type solver_info = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    62
  name: string,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
    63
  class: Proof.context -> SMT_Util.class,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    64
  avail: unit -> bool,
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56078
diff changeset
    65
  options: Proof.context -> string list}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    67
type data = {
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    68
  solvers: (solver_info * string list) Symtab.table,
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    69
  solver: string option,
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    70
  certs: Cache_IO.cache option}
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    71
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    72
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
    73
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    74
val empty_data = mk_data Symtab.empty NONE NONE
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    75
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    76
fun solvers_of ({solvers, ...}: data) = solvers
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    77
fun solver_of ({solver, ...}: data) = solver
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    78
fun certs_of ({certs, ...}: data) = certs
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    79
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    80
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
    81
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
    82
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
    83
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    84
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
    85
  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
    86
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    87
structure Data = Generic_Data
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
(
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    89
  type T = data
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    90
  val empty = empty_data
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    91
  val extend = I
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
    92
  val merge = merge_data
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    93
)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    94
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    95
fun set_solver_options (name, options) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    96
  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
    97
  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
    98
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    99
fun add_solver (info as {name, ...} : solver_info) context =
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   100
  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
   101
    error ("Solver already registered: " ^ quote name)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   102
  else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   103
    context
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   104
    |> Data.map (map_solvers (Symtab.update (name, (info, []))))
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   105
    |> Context.map_theory (Attrib.setup (Binding.name (name ^ "_options"))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   106
        (Scan.lift (@{keyword "="} |-- Args.name) >>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
          (Thm.declaration_attribute o K o set_solver_options o pair name))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   108
        ("Additional command line options for SMT solver " ^ quote name))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   109
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   110
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
   111
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   112
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
   113
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   114
fun is_available ctxt name =
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   115
  (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
   116
    SOME ({avail, ...}, _) => avail ()
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   117
  | NONE => false)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   118
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   119
fun available_solvers_of ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   120
  filter (is_available ctxt) (all_solvers_of ctxt)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   121
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   122
fun warn_solver (Context.Proof ctxt) name =
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56208
diff changeset
   123
      if Context_Position.is_visible ctxt then
57199
blanchet
parents: 57157
diff changeset
   124
        warning ("The SMT solver " ^ quote name ^ " is not installed")
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 56208
diff changeset
   125
      else ()
57199
blanchet
parents: 57157
diff changeset
   126
  | warn_solver _ _ = ()
56078
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 select_solver name context =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   129
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   130
    val ctxt = Context.proof_of context
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   131
    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
   132
  in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   133
    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
   134
      error ("Trying to select unknown solver: " ^ quote name)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   135
    else if not (is_available ctxt name) then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   136
      (warn_solver context name; upd context)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   137
    else upd context
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   138
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   139
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   140
fun no_solver_err () = error "No SMT solver selected"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   141
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   142
fun solver_of ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   143
  (case solver_name_of ctxt of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   144
    SOME name => name
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   145
  | NONE => no_solver_err ())
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   146
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   147
fun solver_info_of default select ctxt =
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   148
  (case solver_name_of ctxt of
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   149
    NONE => default ()
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   150
  | 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
   151
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   152
fun solver_class_of ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   153
  let fun class_of ({class, ...}: solver_info, _) = class ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   154
  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
   155
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   156
fun solver_options_of ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   157
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   158
    fun all_options NONE = []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   159
      | all_options (SOME ({options, ...} : solver_info, opts)) =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   160
          opts @ options ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   161
  in solver_info_of (K []) all_options ctxt end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   162
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   163
val setup_solver =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   164
  Attrib.setup @{binding smt_solver}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
    (Scan.lift (@{keyword "="} |-- Args.name) >>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
      (Thm.declaration_attribute o K o select_solver))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   167
    "SMT solver configuration"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   168
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   169
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   170
(* options *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   171
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   172
val oracle = Attrib.setup_config_bool @{binding smt_oracle} (K true)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   173
val timeout = Attrib.setup_config_real @{binding smt_timeout} (K 30.0)
59213
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
   174
val reconstruction_step_timeout = Attrib.setup_config_real @{binding smt_reconstruction_step_timeout} (K 10.0)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   175
val random_seed = Attrib.setup_config_int @{binding smt_random_seed} (K 1)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   176
val read_only_certificates = Attrib.setup_config_bool @{binding smt_read_only_certificates} (K false)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   177
val verbose = Attrib.setup_config_bool @{binding smt_verbose} (K true)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   178
val trace = Attrib.setup_config_bool @{binding smt_trace} (K false)
59215
35c13f4995b1 optionally display statistics for Z3 proof reconstruction
boehmes
parents: 59214
diff changeset
   179
val statistics = Attrib.setup_config_bool @{binding smt_statistics} (K false)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   180
val monomorph_limit = Attrib.setup_config_int @{binding smt_monomorph_limit} (K 10)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   181
val monomorph_instances = Attrib.setup_config_int @{binding smt_monomorph_instances} (K 500)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   182
val infer_triggers = Attrib.setup_config_bool @{binding smt_infer_triggers} (K false)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   183
val debug_files = Attrib.setup_config_string @{binding smt_debug_files} (K "")
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   184
val sat_solver = Attrib.setup_config_string @{binding smt_sat_solver} (K "cdclite")
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   185
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   186
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   187
(* diagnostics *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   188
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   189
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
   190
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   191
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
   192
fun trace_msg ctxt = cond_trace (Config.get ctxt trace)
59215
35c13f4995b1 optionally display statistics for Z3 proof reconstruction
boehmes
parents: 59214
diff changeset
   193
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
   194
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   195
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   196
(* tools *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   197
59213
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
   198
fun with_time_limit ctxt timeout_config f x =
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
   199
  TimeLimit.timeLimit (seconds (Config.get ctxt timeout_config)) f x
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   200
  handle TimeLimit.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
   201
59213
ef5e68575bc4 limit reconstruction time of Z3 proof steps to be able to detect long-running reconstruction steps
boehmes
parents: 58893
diff changeset
   202
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
   203
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   204
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   205
(* certificates *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   206
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   207
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
   208
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   209
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
   210
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   211
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
   212
  if name = "" then NONE
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   213
  else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   214
    Path.explode name
56208
06cc31dff138 clarifed module name;
wenzelm
parents: 56090
diff changeset
   215
    |> Path.append (Resources.master_directory (Context.theory_of context))
59214
27931bf7720a avoid more than one data slot per module
boehmes
parents: 59213
diff changeset
   216
    |> SOME o Cache_IO.unsynchronized_init))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   217
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   218
val setup_certificates =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57230
diff changeset
   219
  Attrib.setup @{binding smt_certificates}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   220
    (Scan.lift (@{keyword "="} |-- Args.name) >>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   221
      (Thm.declaration_attribute o K o select_certificates))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   222
    "SMT certificates configuration"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   223
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   224
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   225
(* setup *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   226
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   227
val _ = Theory.setup (
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   228
  setup_solver #>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   229
  setup_certificates)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   230
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   231
fun print_setup ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   232
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   233
    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
   234
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   235
    val names = available_solvers_of ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   236
    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
   237
    val n = the_default "(none)" (solver_name_of ctxt)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   238
    val opts = solver_options_of ctxt
57230
ec5ff6bb2a92 eliminate dependency of SMT2 module on 'list'
blanchet
parents: 57229
diff changeset
   239
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   240
    val t = string_of_real (Config.get ctxt timeout)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   241
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   242
    val certs_filename =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   243
      (case get_certificates_path ctxt of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   244
        SOME path => Path.print path
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   245
      | NONE => "(disabled)")
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   246
  in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   247
    Pretty.writeln (Pretty.big_list "SMT setup:" [
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   248
      Pretty.str ("Current SMT solver: " ^ n),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   249
      Pretty.str ("Current SMT solver options: " ^ space_implode " " opts),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   250
      Pretty.str_list "Available SMT solvers: "  "" ns,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   251
      Pretty.str ("Current timeout: " ^ t ^ " seconds"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   252
      Pretty.str ("With proofs: " ^
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   253
        string_of_bool (not (Config.get ctxt oracle))),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   254
      Pretty.str ("Certificates cache: " ^ certs_filename),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   255
      Pretty.str ("Fixed certificates: " ^
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   256
        string_of_bool (Config.get ctxt read_only_certificates))])
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   257
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   258
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   259
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59215
diff changeset
   260
  Outer_Syntax.command @{command_keyword smt_status}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   261
    "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
   262
    \and the values of SMT configuration options"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   263
    (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
   264
57229
blanchet
parents: 57209
diff changeset
   265
end;