src/HOL/SMT/Tools/z3_solver.ML
author boehmes
Tue, 20 Oct 2009 10:11:30 +0200
changeset 33010 39f73a59e855
parent 32618 42865636d006
child 33017 4fb8a33f74d6
permissions -rw-r--r--
added proof reconstructon for Z3, added certificates for simpler re-checking of proofs (no need to invoke external solvers), added examples and certificates for all examples, removed Unsynchronized.ref (in smt_normalize.ML)
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
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    46
    val raw = not o String.isPrefix "WARNING" orf String.isPrefix "ERROR"
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    47
    val (ls, l) = the_default ([], "") (try split_last (filter raw output))
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    48
  in
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    49
    if String.isPrefix "unsat" l then ls
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    50
    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
    51
    else if String.isPrefix "unknown" l then raise_cex false recon ls
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    52
    else error (solver_name ^ " failed")
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    53
  end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    54
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    55
fun core_oracle (SMT_Solver.ProofData {output, recon, ...}) =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    56
  check_unsat recon output
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    57
  |> K @{cprop False}
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    58
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    59
fun prover (SMT_Solver.ProofData {context, output, recon, assms}) =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    60
  check_unsat recon output
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    61
  |> Z3_Proof.reconstruct context assms recon
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    62
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    63
fun solver oracle ctxt =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    64
  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
    65
  in
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    66
    SMT_Solver.SolverConfig {
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents: 32618
diff changeset
    67
      command = {env_var=env_var, remote_name=solver_name},
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents: 32618
diff changeset
    68
      arguments = cmdline_options ctxt,
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    69
      interface = Z3_Interface.interface,
33010
39f73a59e855 added proof reconstructon for Z3,
boehmes
parents: 32618
diff changeset
    70
      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
    71
  end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    72
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    73
val setup =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    74
  proofs_setup #>
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    75
  options_setup #>
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    76
  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
    77
  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
    78
  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
    79
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    80
end