src/HOL/Tools/SMT/smt_solver.ML
author blanchet
Fri, 17 Dec 2010 12:02:46 +0100
changeset 41239 d6e804ff29c3
parent 41131 fc9d503c3d67
child 41241 7d07736aaaf6
permissions -rw-r--r--
split "smt_filter" into head and tail
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Tools/SMT/smt_solver.ML
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     3
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     4
SMT solvers registry and SMT tactic.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     5
*)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     6
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     7
signature SMT_SOLVER =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     8
sig
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
     9
  (*configuration*)
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    10
  datatype outcome = Unsat | Sat | Unknown
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    11
  type solver_config = {
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    12
    name: string,
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    13
    env_var: string,
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    14
    is_remote: bool,
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    15
    options: Proof.context -> string list,
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
    16
    class: SMT_Utils.class,
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    17
    outcome: string -> string list -> outcome * string list,
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    18
    cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
40828
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
    19
      term list * term list) option,
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    20
    reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
    21
      int list * thm) option,
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
    22
    default_max_relevant: int }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    23
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
    24
  (*registry*)
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
    25
  type solver = bool option -> Proof.context ->
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
    26
    (int * (int option * thm)) list -> int list * thm
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    27
  val add_solver: solver_config -> theory -> theory
40940
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
    28
  val solver_name_of: Proof.context -> string
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
    29
  val solver_of: Proof.context -> solver
40940
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
    30
  val available_solvers_of: Proof.context -> string list
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
    31
  val is_locally_installed: Proof.context -> string -> bool
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
    32
  val is_remotely_available: Proof.context -> string -> bool
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
    33
  val default_max_relevant: Proof.context -> string -> int
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    34
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    35
  (*filter*)
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
    36
  type 'a smt_filter_head_result = 'a list * (int option * thm) list *
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
    37
    (((int * thm) list * Proof.context) * (int * (int option * thm)) list)
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
    38
  val smt_filter_head: Time.time -> Proof.state ->
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
    39
    ('a * (int option * thm)) list -> int -> 'a smt_filter_head_result
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
    40
  val smt_filter_tail: bool -> 'a smt_filter_head_result ->
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
    41
    {outcome: SMT_Failure.failure option, used_facts: ('a * thm) list}
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
    42
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    43
  (*tactic*)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    44
  val smt_tac': bool -> Proof.context -> thm list -> int -> Tactical.tactic
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    45
  val smt_tac: Proof.context -> thm list -> int -> Tactical.tactic
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    46
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    47
  (*setup*)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    48
  val setup: theory -> theory
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    49
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    50
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    51
structure SMT_Solver: SMT_SOLVER =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    52
struct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    53
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
    54
structure C = SMT_Config
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    55
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    56
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
    57
(* configuration *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    58
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    59
datatype outcome = Unsat | Sat | Unknown
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    60
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    61
type solver_config = {
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    62
  name: string,
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    63
  env_var: string,
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    64
  is_remote: bool,
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    65
  options: Proof.context -> string list,
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
    66
  class: SMT_Utils.class,
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    67
  outcome: string -> string list -> outcome * string list,
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    68
  cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
40828
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
    69
    term list * term list) option,
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
    70
  reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
    71
    int list * thm) option,
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
    72
  default_max_relevant: int }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    73
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    74
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    75
(* interface to external solvers *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    76
40166
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
    77
fun get_local_solver env_var =
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
    78
  let val local_solver = getenv env_var
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
    79
  in if local_solver <> "" then SOME local_solver else NONE end
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
    80
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    81
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    82
40196
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
    83
fun choose (rm, env_var, is_remote, name) =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    84
  let
40196
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
    85
    val force_local = (case rm of SOME false => true | _ => false)
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
    86
    val force_remote = (case rm of SOME true => true | _ => false)
40166
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
    87
    val lsolver = get_local_solver env_var
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    88
    val remote_url = getenv "REMOTE_SMT_URL"
40196
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
    89
    val trace = if is_some rm then K () else tracing
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    90
  in
40166
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
    91
    if not force_remote andalso is_some lsolver
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    92
    then 
40196
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
    93
     (trace ("Invoking local SMT solver " ^ quote (the lsolver) ^ " ...");
40166
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
    94
      [the lsolver])
40196
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
    95
    else if not force_local andalso is_remote
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    96
    then
40196
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
    97
     (trace ("Invoking remote SMT solver " ^ quote name ^ " at " ^
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    98
        quote remote_url ^ " ...");
40166
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
    99
      [getenv "REMOTE_SMT", name])
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
   100
    else if force_remote
40196
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
   101
    then error ("The SMT solver " ^ quote name ^ " is not remotely available.")
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
   102
    else error ("The SMT solver " ^ quote name ^ " has not been found " ^
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
   103
      "on this computer. Please set the Isabelle environment variable " ^
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
   104
      quote env_var ^ ".")
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   105
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   106
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   107
fun make_cmd solver args problem_path proof_path = space_implode " " (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   108
  map File.shell_quote (solver @ args) @
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   109
  [File.shell_path problem_path, "2>&1", ">", File.shell_path proof_path])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   110
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   111
fun run ctxt cmd args input =
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   112
  (case C.certificates_of ctxt of
40578
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   113
    NONE =>
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   114
      if Config.get ctxt C.debug_files = "" then
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   115
        Cache_IO.run (make_cmd (choose cmd) args) input
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   116
      else
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   117
        let
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   118
          val base_path = Path.explode (Config.get ctxt C.debug_files)
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   119
          val in_path = Path.ext "smt_in" base_path
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   120
          val out_path = Path.ext "smt_out" base_path
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   121
        in
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   122
          Cache_IO.raw_run (make_cmd (choose cmd) args) input in_path out_path 
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
   123
        end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   124
  | SOME certs =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   125
      (case Cache_IO.lookup certs input of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   126
        (NONE, key) =>
40538
b8482ff0bc92 check the return code of the SMT solver and raise an exception if the prover failed
boehmes
parents: 40515
diff changeset
   127
          if Config.get ctxt C.fixed then
b8482ff0bc92 check the return code of the SMT solver and raise an exception if the prover failed
boehmes
parents: 40515
diff changeset
   128
            error ("Bad certificates cache: missing certificate")
b8482ff0bc92 check the return code of the SMT solver and raise an exception if the prover failed
boehmes
parents: 40515
diff changeset
   129
          else
b8482ff0bc92 check the return code of the SMT solver and raise an exception if the prover failed
boehmes
parents: 40515
diff changeset
   130
            Cache_IO.run_and_cache certs key (make_cmd (choose cmd) args) input
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   131
      | (SOME output, _) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   132
         (tracing ("Using cached certificate from " ^
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   133
            File.shell_path (Cache_IO.cache_path_of certs) ^ " ...");
40550
f84c664ece8e trace more solver output before raising an exception due to a non-zero return code (avoids truncating potential counterexamples produced by Z3)
boehmes
parents: 40538
diff changeset
   134
          output)))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   135
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   136
in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   137
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   138
fun run_solver ctxt cmd args input =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   139
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   140
    fun pretty tag ls = Pretty.string_of (Pretty.big_list tag
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   141
      (map Pretty.str ls))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   142
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   143
    val _ = C.trace_msg ctxt (pretty "Problem:" o split_lines) input
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   144
40550
f84c664ece8e trace more solver output before raising an exception due to a non-zero return code (avoids truncating potential counterexamples produced by Z3)
boehmes
parents: 40538
diff changeset
   145
    val {redirected_output=res, output=err, return_code} =
f84c664ece8e trace more solver output before raising an exception due to a non-zero return code (avoids truncating potential counterexamples produced by Z3)
boehmes
parents: 40538
diff changeset
   146
      C.with_timeout ctxt (run ctxt cmd args) input
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   147
    val _ = C.trace_msg ctxt (pretty "Solver:") err
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   148
39811
0659e84bdc5f chop_while replace drop_while and take_while
haftmann
parents: 39809
diff changeset
   149
    val ls = rev (snd (chop_while (equal "") (rev res)))
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   150
    val _ = C.trace_msg ctxt (pretty "Result:") ls
40550
f84c664ece8e trace more solver output before raising an exception due to a non-zero return code (avoids truncating potential counterexamples produced by Z3)
boehmes
parents: 40538
diff changeset
   151
f84c664ece8e trace more solver output before raising an exception due to a non-zero return code (avoids truncating potential counterexamples produced by Z3)
boehmes
parents: 40538
diff changeset
   152
    val _ = null ls andalso return_code <> 0 andalso
40561
0125cbb5d3c7 renamed SMT failure: Abnormal_Termination is indeed more appropriate than Solver_Crashed
boehmes
parents: 40560
diff changeset
   153
      raise SMT_Failure.SMT (SMT_Failure.Abnormal_Termination return_code)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   154
  in ls end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   155
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   156
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   157
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   158
fun trace_assms ctxt = C.trace_msg ctxt (Pretty.string_of o
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   159
  Pretty.big_list "Assertions:" o map (Display.pretty_thm ctxt o snd))
40198
8d470bbaafd7 trace assumptions before giving them to the SMT solver
boehmes
parents: 40197
diff changeset
   160
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   161
fun trace_recon_data ({context=ctxt, typs, terms, ...} : SMT_Translate.recon) =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   162
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   163
    fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   164
    fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   165
    fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   166
  in
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   167
    C.trace_msg ctxt (fn () =>
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   168
      Pretty.string_of (Pretty.big_list "Names:" [
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   169
        Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)),
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   170
        Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   171
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   172
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   173
fun invoke name cmd options ithms ctxt =
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   174
  let
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   175
    val args = C.solver_options_of ctxt @ options ctxt
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   176
    val comments = ("solver: " ^ name) ::
41062
boehmes
parents: 41059
diff changeset
   177
      ("timeout: " ^ string_of_real (Config.get ctxt C.timeout)) ::
41121
5c5d05963f93 added option to modify the random seed of SMT solvers
boehmes
parents: 41065
diff changeset
   178
      ("random seed: " ^ string_of_int (Config.get ctxt C.random_seed)) ::
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   179
      "arguments:" :: args
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   180
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   181
    val (str, recon as {context=ctxt', ...}) =
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   182
      ithms
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   183
      |> tap (trace_assms ctxt)
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   184
      |> SMT_Translate.translate ctxt comments
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   185
      ||> tap trace_recon_data
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   186
  in (run_solver ctxt' cmd args str, recon) end
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   187
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   188
fun trace_assumptions ctxt iwthms idxs =
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   189
  let
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   190
    val wthms =
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   191
      idxs
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   192
      |> filter (fn i => i >= 0)
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   193
      |> map_filter (AList.lookup (op =) iwthms)
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   194
  in
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   195
    if Config.get ctxt C.trace_used_facts andalso length wthms > 0
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   196
    then
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   197
      tracing (Pretty.string_of (Pretty.big_list "SMT used facts:"
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   198
        (map (Display.pretty_thm ctxt o snd) wthms)))
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   199
    else ()
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   200
  end
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   201
41041
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   202
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   203
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   204
(* registry *)
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   205
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   206
type solver = bool option -> Proof.context -> (int * (int option * thm)) list ->
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   207
  int list * thm
41041
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   208
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   209
type solver_info = {
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   210
  env_var: string,
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   211
  is_remote: bool,
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   212
  options: Proof.context -> string list,
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   213
  reconstruct: Proof.context -> string list * SMT_Translate.recon ->
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   214
    int list * thm,
41041
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   215
  default_max_relevant: int }
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   216
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   217
fun gen_solver_head ctxt iwthms =
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   218
  SMT_Normalize.normalize ctxt iwthms
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   219
  |> rpair ctxt
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   220
  |-> SMT_Monomorph.monomorph
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   221
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   222
fun gen_solver_tail (name, info : solver_info) rm (iwthms', ctxt) iwthms =
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   223
  let
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   224
    val {env_var, is_remote, options, reconstruct, ...} = info
40196
123b6fe379f6 honor choice of either local or remote SMT solver only for smt_filter and keep default behaviour for the SMT tactic; omit messages when running smt_filter
boehmes
parents: 40166
diff changeset
   225
    val cmd = (rm, env_var, is_remote, name)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   226
  in
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   227
    (iwthms', ctxt)
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   228
    |-> invoke name cmd options
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   229
    |> reconstruct ctxt
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   230
    |> (fn (idxs, thm) => thm
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   231
    |> tap (fn _ => trace_assumptions ctxt iwthms idxs)
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   232
    |> pair idxs)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   233
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   234
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   235
structure Solvers = Generic_Data
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   236
(
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   237
  type T = solver_info Symtab.table
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   238
  val empty = Symtab.empty
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   239
  val extend = I
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   240
  fun merge data = Symtab.merge (K true) data
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   241
    handle Symtab.DUP name => error ("Duplicate SMT solver: " ^ quote name)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   242
)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   243
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   244
local
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   245
  fun finish outcome cex_parser reconstruct ocl outer_ctxt
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   246
      (output, (recon as {context=ctxt, ...} : SMT_Translate.recon)) =
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   247
    (case outcome output of
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   248
      (Unsat, ls) =>
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   249
        if not (Config.get ctxt C.oracle) andalso is_some reconstruct
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   250
        then the reconstruct outer_ctxt recon ls
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   251
        else ([], ocl ())
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   252
    | (result, ls) =>
40828
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   253
        let
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   254
          val (ts, us) =
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   255
            (case cex_parser of SOME f => f ctxt recon ls | _ => ([], []))
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   256
         in
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   257
          raise SMT_Failure.SMT (SMT_Failure.Counterexample {
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   258
            is_real_cex = (result = Sat),
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   259
            free_constraints = ts,
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   260
            const_defs = us})
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   261
        end)
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40578
diff changeset
   262
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40578
diff changeset
   263
  val cfalse = Thm.cterm_of @{theory} (@{const Trueprop} $ @{const False})
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   264
in
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   265
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   266
fun add_solver cfg =
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   267
  let
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   268
    val {name, env_var, is_remote, options, class, outcome, cex_parser,
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   269
      reconstruct, default_max_relevant} = cfg
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   270
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40578
diff changeset
   271
    fun core_oracle () = cfalse
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   272
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   273
    fun solver ocl = { env_var=env_var, is_remote=is_remote, options=options,
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   274
      reconstruct=finish (outcome name) cex_parser reconstruct ocl,
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   275
      default_max_relevant=default_max_relevant }
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   276
  in
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   277
    Thm.add_oracle (Binding.name name, core_oracle) #-> (fn (_, ocl) =>
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   278
    Context.theory_map (Solvers.map (Symtab.update_new (name, solver ocl)))) #>
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 40981
diff changeset
   279
    Context.theory_map (C.add_solver (name, class))
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   280
  end
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   281
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   282
end
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   283
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   284
fun get_info ctxt name =
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   285
  the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   286
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   287
fun name_and_solver_of ctxt =
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   288
  let val name = C.solver_of ctxt
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   289
  in (name, get_info ctxt name) end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   290
40940
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
   291
val solver_name_of = fst o name_and_solver_of
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   292
fun solver_of ctxt rm ctxt' =
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   293
  `(gen_solver_head ctxt') #-> gen_solver_tail (name_and_solver_of ctxt) rm
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   294
40940
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
   295
val available_solvers_of = Symtab.keys o Solvers.get o Context.Proof
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
   296
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
   297
fun is_locally_installed ctxt name =
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   298
  let val {env_var, ...} = get_info ctxt name
40166
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
   299
  in is_some (get_local_solver env_var) end
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
   300
40981
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   301
val is_remotely_available = #is_remote oo get_info
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   302
67f436af0638 have SMT solvers report the number of facts that they should have by default in Sledgehammer -- the information might not seem to belong there but it also belongs nowhere else, for how is Sledgehammer to know how different solvers deal with hundreds of facts?
blanchet
parents: 40980
diff changeset
   303
val default_max_relevant = #default_max_relevant oo get_info
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   304
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   305
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   306
(* filter *)
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   307
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   308
val has_topsort = Term.exists_type (Term.exists_subtype (fn
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   309
    TFree (_, []) => true
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   310
  | TVar (_, []) => true
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   311
  | _ => false))
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   312
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   313
(* without this test, we would run into problems when atomizing the rules: *)
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   314
fun check_topsort iwthms =
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   315
  if exists (has_topsort o Thm.prop_of o snd o snd) iwthms then
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   316
    raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("proof state " ^
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   317
      "contains the universal sort {}"))
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   318
  else
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   319
    ()
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   320
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40578
diff changeset
   321
val cnot = Thm.cterm_of @{theory} @{const Not}
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40578
diff changeset
   322
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   323
fun mk_result outcome xrules = { outcome = outcome, used_facts = xrules }
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   324
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   325
type 'a smt_filter_head_result = 'a list * (int option * thm) list *
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   326
  (((int * thm) list * Proof.context) * (int * (int option * thm)) list)
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   327
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   328
fun smt_filter_head time_limit st xwrules i =
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   329
  let
40164
57f5db2a48a3 added a mode to only filter assumptions used in a Z3 proof (in which no proof reconstruction is performed)
boehmes
parents: 40162
diff changeset
   330
    val ctxt =
40199
2963511e121c use proper context
boehmes
parents: 40198
diff changeset
   331
      Proof.context_of st
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   332
      |> Config.put C.oracle false
40560
b57f7fee72ee honour timeouts which are not rounded to full seconds
boehmes
parents: 40550
diff changeset
   333
      |> Config.put C.timeout (Time.toReal time_limit)
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   334
      |> Config.put C.drop_bad_facts true
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   335
      |> Config.put C.filter_only_facts true
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   336
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   337
    val {facts, goal, ...} = Proof.goal st
40357
82ebdd19c4a4 simulate more closely the behaviour of the tactic
boehmes
parents: 40332
diff changeset
   338
    val ({context=ctxt', prems, concl, ...}, _) = Subgoal.focus ctxt i goal
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40578
diff changeset
   339
    fun negate ct = Thm.dest_comb ct ||> Thm.capply cnot |-> Thm.capply
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40578
diff changeset
   340
    val cprop = negate (Thm.rhs_of (SMT_Normalize.atomize_conv ctxt' concl))
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   341
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   342
    val (xs, wthms) = split_list xwrules
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   343
  in
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   344
    (xs, wthms,
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   345
     wthms
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   346
     |> map_index I
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   347
     |> append (map (pair ~1 o pair NONE) (Thm.assume cprop :: prems @ facts))
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   348
     |> tap check_topsort
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   349
     |> `(gen_solver_head ctxt'))
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   350
  end
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   351
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   352
fun smt_filter_tail run_remote (xs, wthms, head as ((_, ctxt), _)) =
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   353
  let val xrules = xs ~~ map snd wthms in
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   354
    head
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   355
    |-> gen_solver_tail (name_and_solver_of ctxt) (SOME run_remote)
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   356
    |> distinct (op =) o fst
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   357
    |> map_filter (try (nth xrules))
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   358
    |> (if solver_name_of ctxt = "z3" (* FIXME *) then I else K xrules)
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   359
    |> mk_result NONE
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   360
  end
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   361
  handle SMT_Failure.SMT fail => mk_result (SOME fail) []
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   362
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39811
diff changeset
   363
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   364
(* SMT tactic *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   365
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   366
fun smt_tac' pass_exns ctxt rules =
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   367
  CONVERSION (SMT_Normalize.atomize_conv ctxt)
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   368
  THEN' Tactic.rtac @{thm ccontr}
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   369
  THEN' SUBPROOF (fn {context=ctxt', prems, ...} =>
40165
boehmes
parents: 40164
diff changeset
   370
    let
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   371
      val solve = snd o solver_of ctxt' NONE ctxt' o tap check_topsort
40424
7550b2cba1cb better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents: 40357
diff changeset
   372
      val tag = "Solver " ^ C.solver_of ctxt' ^ ": "
40828
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   373
      val str_of = prefix tag o SMT_Failure.string_of_failure ctxt'
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   374
      fun safe_solve iwthms =
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   375
        if pass_exns then SOME (solve iwthms)
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   376
        else (SOME (solve iwthms)
40515
25f266144206 turned SMT counterexamples into verbose messages (they had been swallowed before, following the state of smt_trace -- which is off by default), because they might be useful for the user
boehmes
parents: 40425
diff changeset
   377
          handle
25f266144206 turned SMT counterexamples into verbose messages (they had been swallowed before, following the state of smt_trace -- which is off by default), because they might be useful for the user
boehmes
parents: 40425
diff changeset
   378
            SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) =>
40828
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   379
              (C.verbose_msg ctxt' str_of fail; NONE)
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   380
          | SMT_Failure.SMT fail => (C.trace_msg ctxt' str_of fail; NONE))
40165
boehmes
parents: 40164
diff changeset
   381
    in
41126
e0bd443c0fdd re-ordered SMT normalization code (eta-normalization, lambda abstractions and partial functions will be dealt with on the term level);
boehmes
parents: 41124
diff changeset
   382
      safe_solve (map (pair ~1 o pair NONE) (rules @ prems))
40165
boehmes
parents: 40164
diff changeset
   383
      |> (fn SOME thm => Tactic.rtac thm 1 | _ => Tactical.no_tac)
40166
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
   384
    end) ctxt
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   385
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   386
val smt_tac = smt_tac' false
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   387
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   388
val smt_method =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   389
  Scan.optional Attrib.thms [] >>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   390
  (fn thms => fn ctxt => METHOD (fn facts =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   391
    HEADGOAL (smt_tac ctxt (thms @ facts))))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   392
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   393
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   394
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   395
(* setup *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   396
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   397
val setup =
38808
89ae86205739 more antiquotations;
wenzelm
parents: 36960
diff changeset
   398
  Method.setup @{binding smt} smt_method
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   399
    "Applies an SMT solver to the current goal."
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   400
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   401
end