src/HOL/SMT/Tools/z3_solver.ML
changeset 36898 8e55aa1306c5
parent 36897 6d1ecdb81ff0
child 36899 bcd6fce5bf06
equal deleted inserted replaced
36897:6d1ecdb81ff0 36898:8e55aa1306c5
     1 (*  Title:      HOL/SMT/Tools/z3_solver.ML
       
     2     Author:     Sascha Boehme, TU Muenchen
       
     3 
       
     4 Interface of the SMT solver Z3.
       
     5 *)
       
     6 
       
     7 signature Z3_SOLVER =
       
     8 sig
       
     9   val proofs: bool Config.T
       
    10   val options: string Config.T
       
    11   val setup: theory -> theory
       
    12 end
       
    13 
       
    14 structure Z3_Solver: Z3_SOLVER =
       
    15 struct
       
    16 
       
    17 val solver_name = "z3"
       
    18 val env_var = "Z3_SOLVER"
       
    19 
       
    20 val (proofs, proofs_setup) = Attrib.config_bool "z3_proofs" (K false)
       
    21 val (options, options_setup) = Attrib.config_string "z3_options" (K "")
       
    22 
       
    23 fun add xs ys = ys @ xs
       
    24 
       
    25 fun explode_options s = String.tokens (Symbol.is_ascii_blank o str) s
       
    26 
       
    27 fun get_options ctxt =
       
    28   ["MODEL=true", "PRE_SIMPLIFY_EXPR=false", "CONTEXT_SIMPLIFIER=false"]
       
    29   |> Config.get ctxt proofs ? add ["DISPLAY_PROOF=true", "PROOF_MODE=2"]
       
    30   |> add (explode_options (Config.get ctxt options))
       
    31 
       
    32 fun pretty_config context = [
       
    33   Pretty.str ("With proofs: " ^
       
    34     (if Config.get_generic context proofs then "true" else "false")),
       
    35   Pretty.str ("Options: " ^
       
    36     space_implode " " (get_options (Context.proof_of context))) ]
       
    37 
       
    38 fun cmdline_options ctxt =
       
    39   get_options ctxt
       
    40   |> add ["-smt"]
       
    41 
       
    42 fun raise_cex real recon ls =
       
    43   let val cex = Z3_Model.parse_counterex recon ls
       
    44   in raise SMT_Solver.SMT_COUNTEREXAMPLE (real, cex) end
       
    45 
       
    46 fun if_unsat f (output, recon) =
       
    47   let
       
    48     fun jnk l =
       
    49       String.isPrefix "WARNING" l orelse
       
    50       String.isPrefix "ERROR" l orelse
       
    51       forall Symbol.is_ascii_blank (Symbol.explode l)
       
    52     val (ls, l) = the_default ([], "") (try split_last (filter_out jnk output))
       
    53   in
       
    54     if String.isPrefix "unsat" l then f (ls, recon)
       
    55     else if String.isPrefix "sat" l then raise_cex true recon ls
       
    56     else if String.isPrefix "unknown" l then raise_cex false recon ls
       
    57     else raise SMT_Solver.SMT (solver_name ^ " failed")
       
    58   end
       
    59 
       
    60 val core_oracle = if_unsat (K @{cprop False})
       
    61 
       
    62 val prover = if_unsat Z3_Proof_Reconstruction.reconstruct
       
    63 
       
    64 fun solver oracle ctxt =
       
    65   let val with_proof = Config.get ctxt proofs
       
    66   in
       
    67    {command = {env_var=env_var, remote_name=SOME solver_name},
       
    68     arguments = cmdline_options ctxt,
       
    69     interface = Z3_Interface.interface,
       
    70     reconstruct = if with_proof then prover else pair o oracle}
       
    71   end
       
    72 
       
    73 val setup =
       
    74   proofs_setup #>
       
    75   options_setup #>
       
    76   Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) =>
       
    77   SMT_Solver.add_solver (solver_name, solver oracle)) #>
       
    78   SMT_Solver.add_solver_info (solver_name, pretty_config)
       
    79 
       
    80 end