src/HOL/SMT/Tools/z3_solver.ML
author boehmes
Tue, 03 Nov 2009 14:51:55 +0100
changeset 33418 1312e8337ce5
parent 33299 73af7831ba1e
child 33472 e88f67d679c4
permissions -rw-r--r--
added a specific SMT exception captured by smt_tac (prevents the SMT method from failing with an exception), replaced unspecific 'error' invocations with raising of specific SMT exceptions, added annotations to traced SMT problem and solver output
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/SMT/Tools/z3_solver.ML
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     3
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     4
Interface of the SMT solver Z3.
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     5
*)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     6
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     7
signature Z3_SOLVER =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     8
sig
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     9
  val proofs: bool Config.T
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    10
  val options: string Config.T
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    11
  val setup: theory -> theory
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    12
end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    13
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    14
structure Z3_Solver: Z3_SOLVER =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    15
struct
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    16
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    17
val solver_name = "z3"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    18
val env_var = "Z3_SOLVER"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    19
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    20
val (proofs, proofs_setup) = Attrib.config_bool "z3_proofs" false
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    21
val (options, options_setup) = Attrib.config_string "z3_options" ""
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    22
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    23
fun add xs ys = ys @ xs
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    24
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    25
fun get_options ctxt =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    26
  ["MODEL=true", "PRE_SIMPLIFY_EXPR=false", "CONTEXT_SIMPLIFIER=false"]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    27
  |> Config.get ctxt proofs ? add ["DISPLAY_PROOF=true", "PROOF_MODE=2"]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    28
  |> add (space_explode " " (Config.get ctxt options))
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    29
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    30
fun pretty_config context = [
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    31
  Pretty.str ("With proofs: " ^
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    32
    (if Config.get_generic context proofs then "true" else "false")),
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    33
  Pretty.str ("Options: " ^
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    34
    space_implode " " (get_options (Context.proof_of context))) ]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    35
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    36
fun cmdline_options ctxt =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    37
  get_options ctxt
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    38
  |> add ["-smt"]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    39
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    40
fun raise_cex real recon ls =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    41
  let val cex = Z3_Model.parse_counterex recon ls
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    42
  in raise SMT_Solver.SMT_COUNTEREXAMPLE (real, cex) end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    43
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    44
fun check_unsat recon output =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    45
  let
33299
73af7831ba1e simplified method syntax of "smt",
boehmes
parents: 33017
diff changeset
    46
    fun jnk l =
73af7831ba1e simplified method syntax of "smt",
boehmes
parents: 33017
diff changeset
    47
      String.isPrefix "WARNING" l orelse
73af7831ba1e simplified method syntax of "smt",
boehmes
parents: 33017
diff changeset
    48
      String.isPrefix "ERROR" l orelse
73af7831ba1e simplified method syntax of "smt",
boehmes
parents: 33017
diff changeset
    49
      forall Symbol.is_ascii_blank (Symbol.explode l)
73af7831ba1e simplified method syntax of "smt",
boehmes
parents: 33017
diff changeset
    50
    val (ls, l) = the_default ([], "") (try split_last (filter_out jnk output))
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    51
  in
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    52
    if String.isPrefix "unsat" l then ls
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    53
    else if String.isPrefix "sat" l then raise_cex true recon ls
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    54
    else if String.isPrefix "unknown" l then raise_cex false recon ls
33418
1312e8337ce5 added a specific SMT exception captured by smt_tac (prevents the SMT method from failing with an exception),
boehmes
parents: 33299
diff changeset
    55
    else raise SMT_Solver.SMT (solver_name ^ " failed")
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    56
  end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    57
33017
4fb8a33f74d6 eliminated extraneous wrapping of public records,
boehmes
parents: 33010
diff changeset
    58
fun core_oracle ({output, recon, ...} : SMT_Solver.proof_data) =
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    59
  check_unsat recon output
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    60
  |> K @{cprop False}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    61
33017
4fb8a33f74d6 eliminated extraneous wrapping of public records,
boehmes
parents: 33010
diff changeset
    62
fun prover ({context, output, recon, assms} : SMT_Solver.proof_data) =
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    63
  check_unsat recon output
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    64
  |> Z3_Proof.reconstruct context assms recon
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    65
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    66
fun solver oracle ctxt =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    67
  let val with_proof = Config.get ctxt proofs
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    68
  in
33017
4fb8a33f74d6 eliminated extraneous wrapping of public records,
boehmes
parents: 33010
diff changeset
    69
    {command = {env_var=env_var, remote_name=solver_name},
4fb8a33f74d6 eliminated extraneous wrapping of public records,
boehmes
parents: 33010
diff changeset
    70
    arguments = cmdline_options ctxt,
4fb8a33f74d6 eliminated extraneous wrapping of public records,
boehmes
parents: 33010
diff changeset
    71
    interface = Z3_Interface.interface,
4fb8a33f74d6 eliminated extraneous wrapping of public records,
boehmes
parents: 33010
diff changeset
    72
    reconstruct = if with_proof then prover else oracle}
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    73
  end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    74
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    75
val setup =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    76
  proofs_setup #>
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    77
  options_setup #>
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    78
  Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) =>
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    79
  SMT_Solver.add_solver (solver_name, solver oracle)) #>
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    80
  SMT_Solver.add_solver_info (solver_name, pretty_config)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    81
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    82
end