src/HOL/Tools/SMT/smt_solver.ML
author blanchet
Tue, 27 Mar 2012 16:59:13 +0300
changeset 47152 446cfc760ccf
parent 46743 8e365bc843e9
child 47701 157e6108a342
permissions -rw-r--r--
renamed "smt_fixed" to "smt_read_only_certificates"
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,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    13
    class: SMT_Utils.class,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    14
    avail: unit -> bool,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    15
    command: unit -> string list,
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
    16
    options: Proof.context -> string list,
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    17
    default_max_relevant: int,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    18
    supports_filter: bool,
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
    19
    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
    20
    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
    21
      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
    22
    reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    23
      int list * thm) option }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    24
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
    25
  (*registry*)
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
    26
  val add_solver: solver_config -> theory -> theory
40940
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
    27
  val solver_name_of: Proof.context -> string
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
    28
  val available_solvers_of: Proof.context -> string list
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    29
  val apply_solver: Proof.context -> (int * (int option * thm)) list ->
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    30
    int list * thm
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
    31
  val default_max_relevant: Proof.context -> string -> int
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    32
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
    33
  (*filter*)
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    34
  type 'a smt_filter_data =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    35
    ('a * thm) list * ((int * thm) list * Proof.context)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    36
  val smt_filter_preprocess: Proof.state ->
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    37
    ('a * (int option * thm)) list -> int -> 'a smt_filter_data
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    38
  val smt_filter_apply: Time.time -> 'a smt_filter_data ->
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
    39
    {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
    40
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    41
  (*tactic*)
46464
4cf5a84e2c05 normalized aliases;
wenzelm
parents: 42616
diff changeset
    42
  val smt_tac: Proof.context -> thm list -> int -> tactic
4cf5a84e2c05 normalized aliases;
wenzelm
parents: 42616
diff changeset
    43
  val smt_tac': Proof.context -> thm list -> int -> tactic
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    44
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    45
  (*setup*)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    46
  val setup: theory -> theory
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    47
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    48
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    49
structure SMT_Solver: SMT_SOLVER =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    50
struct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    51
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
    52
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    53
(* interface to external solvers *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    54
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    55
local
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    56
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    57
fun make_cmd command options problem_path proof_path = space_implode " " (
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41497
diff changeset
    58
  map File.shell_quote (command () @ options) @
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    59
  [File.shell_path problem_path, "2>&1", ">", File.shell_path proof_path])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    60
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    61
fun trace_and ctxt msg f x =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    62
  let val _ = SMT_Config.trace_msg ctxt (fn () => msg) ()
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    63
  in f x end
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    64
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    65
fun run ctxt name mk_cmd input =
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
    66
  (case SMT_Config.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
    67
    NONE =>
46743
8e365bc843e9 fine-tune intended failure of smt method when the chosen SMT solver is not installed: restrict failures to true invocations of the SMT solver and don't fail for runs based on certificates
boehmes
parents: 46736
diff changeset
    68
      if not (SMT_Config.is_available ctxt name) then
8e365bc843e9 fine-tune intended failure of smt method when the chosen SMT solver is not installed: restrict failures to true invocations of the SMT solver and don't fail for runs based on certificates
boehmes
parents: 46736
diff changeset
    69
        error ("The SMT solver " ^ quote name ^ " is not installed.")
8e365bc843e9 fine-tune intended failure of smt method when the chosen SMT solver is not installed: restrict failures to true invocations of the SMT solver and don't fail for runs based on certificates
boehmes
parents: 46736
diff changeset
    70
      else if Config.get ctxt SMT_Config.debug_files = "" then
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    71
        trace_and ctxt ("Invoking SMT solver " ^ quote name ^ " ...")
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    72
          (Cache_IO.run mk_cmd) input
40578
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
    73
      else
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
    74
        let
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
    75
          val base_path = Path.explode (Config.get ctxt SMT_Config.debug_files)
40578
2b098a549450 keep input and output files used to communicate with the SMT solver (for debugging purposes mainly)
boehmes
parents: 40561
diff changeset
    76
          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
    77
          val out_path = Path.ext "smt_out" base_path
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    78
        in Cache_IO.raw_run mk_cmd input in_path out_path end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    79
  | SOME certs =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    80
      (case Cache_IO.lookup certs input of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    81
        (NONE, key) =>
47152
446cfc760ccf renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents: 46743
diff changeset
    82
          if Config.get ctxt SMT_Config.read_only_certificates then
446cfc760ccf renamed "smt_fixed" to "smt_read_only_certificates"
blanchet
parents: 46743
diff changeset
    83
            error ("Bad certificate cache: missing certificate")
40538
b8482ff0bc92 check the return code of the SMT solver and raise an exception if the prover failed
boehmes
parents: 40515
diff changeset
    84
          else
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    85
            Cache_IO.run_and_cache certs key mk_cmd input
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    86
      | (SOME output, _) =>
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    87
          trace_and ctxt ("Using cached certificate from " ^
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    88
            File.shell_path (Cache_IO.cache_path_of certs) ^ " ...")
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    89
            I output))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    90
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    91
fun run_solver ctxt name mk_cmd input =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    92
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    93
    fun pretty tag ls = Pretty.string_of (Pretty.big_list tag
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    94
      (map Pretty.str ls))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    95
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
    96
    val _ = SMT_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    97
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
    98
    val {redirected_output=res, output=err, return_code} =
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
    99
      SMT_Config.with_timeout ctxt (run ctxt name mk_cmd) input
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   100
    val _ = SMT_Config.trace_msg ctxt (pretty "Solver:") err
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   101
39811
0659e84bdc5f chop_while replace drop_while and take_while
haftmann
parents: 39809
diff changeset
   102
    val ls = rev (snd (chop_while (equal "") (rev res)))
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   103
    val _ = SMT_Config.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
   104
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
   105
    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
   106
      raise SMT_Failure.SMT (SMT_Failure.Abnormal_Termination return_code)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   107
  in ls end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   108
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   109
fun trace_assms ctxt =
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   110
  SMT_Config.trace_msg ctxt (Pretty.string_of o
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   111
    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
   112
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
   113
fun trace_recon_data ({context=ctxt, typs, terms, ...} : SMT_Translate.recon) =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   114
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   115
    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
   116
    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
   117
    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
   118
  in
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   119
    SMT_Config.trace_msg ctxt (fn () =>
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
   120
      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
   121
        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
   122
        Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   123
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   124
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   125
in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   126
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   127
fun invoke name command 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
   128
  let
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   129
    val options = SMT_Config.solver_options_of 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
   130
    val comments = ("solver: " ^ name) ::
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   131
      ("timeout: " ^ string_of_real (Config.get ctxt SMT_Config.timeout)) ::
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   132
      ("random seed: " ^
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   133
        string_of_int (Config.get ctxt SMT_Config.random_seed)) ::
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   134
      "arguments:" :: options
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   135
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
   136
    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
   137
      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
   138
      |> 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
   139
      |> 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
   140
      ||> tap trace_recon_data
41601
fda8511006f9 made Z3 the default SMT solver again
boehmes
parents: 41497
diff changeset
   141
  in (run_solver ctxt' name (make_cmd command options) str, recon) end
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   142
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   143
end
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   144
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   145
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   146
(* configuration *)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   147
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   148
datatype outcome = Unsat | Sat | Unknown
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
   149
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   150
type solver_config = {
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   151
  name: string,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   152
  class: SMT_Utils.class,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   153
  avail: unit -> bool,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   154
  command: unit -> string list,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   155
  options: Proof.context -> string list,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   156
  default_max_relevant: int,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   157
  supports_filter: bool,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   158
  outcome: string -> string list -> outcome * string list,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   159
  cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   160
    term list * term list) option,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   161
  reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   162
    int list * thm) option }
41041
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   163
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   164
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   165
(* registry *)
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   166
ec2734f34d0f make SML/NJ happy
blanchet
parents: 40981
diff changeset
   167
type solver_info = {
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   168
  command: unit -> string list,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   169
  default_max_relevant: int,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   170
  supports_filter: bool,
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
   171
  reconstruct: Proof.context -> string list * SMT_Translate.recon ->
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   172
    int list * thm }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   173
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   174
structure Solvers = Generic_Data
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   175
(
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
  type T = solver_info Symtab.table
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   177
  val empty = Symtab.empty
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   178
  val extend = I
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   179
  fun merge data = Symtab.merge (K true) data
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   180
)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   181
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
   182
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
   183
  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
   184
      (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
   185
    (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
   186
      (Unsat, ls) =>
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   187
        if not (Config.get ctxt SMT_Config.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
   188
        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
   189
        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
   190
    | (result, ls) =>
40828
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   191
        let
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   192
          val (ts, us) =
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   193
            (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
   194
         in
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   195
          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
   196
            is_real_cex = (result = Sat),
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   197
            free_constraints = ts,
47ff261431c4 split up Z3 models into constraints on free variables and constant definitions;
boehmes
parents: 40666
diff changeset
   198
            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
   199
        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
   200
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
   201
  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
   202
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
   203
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
   204
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
   205
  let
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   206
    val {name, class, avail, command, options, default_max_relevant,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   207
      supports_filter, outcome, cex_parser, reconstruct} = 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
   208
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
   209
    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
   210
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   211
    fun solver ocl = {
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   212
      command = command,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   213
      default_max_relevant = default_max_relevant,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   214
      supports_filter = supports_filter,
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   215
      reconstruct = finish (outcome name) cex_parser reconstruct ocl }
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   216
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   217
    val info = {name=name, class=class, avail=avail, options=options}
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
   218
  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
   219
    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
   220
    Context.theory_map (Solvers.map (Symtab.update_new (name, solver ocl)))) #>
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   221
    Context.theory_map (SMT_Config.add_solver info)
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
   222
  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
   223
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
   224
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
   225
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
   226
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
   227
  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
   228
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   229
val solver_name_of = SMT_Config.solver_of
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   230
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   231
val available_solvers_of = SMT_Config.available_solvers_of
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   232
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   233
fun name_and_info_of ctxt =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   234
  let val name = solver_name_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
   235
  in (name, get_info ctxt name) end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   236
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   237
fun gen_preprocess ctxt iwthms = SMT_Normalize.normalize iwthms ctxt
40940
ff805bb109d8 export more information about available SMT solvers
blanchet
parents: 40828
diff changeset
   238
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   239
fun gen_apply (ithms, ctxt) =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   240
  let val (name, {command, reconstruct, ...}) = name_and_info_of ctxt
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   241
  in
46743
8e365bc843e9 fine-tune intended failure of smt method when the chosen SMT solver is not installed: restrict failures to true invocations of the SMT solver and don't fail for runs based on certificates
boehmes
parents: 46736
diff changeset
   242
    (ithms, ctxt)
8e365bc843e9 fine-tune intended failure of smt method when the chosen SMT solver is not installed: restrict failures to true invocations of the SMT solver and don't fail for runs based on certificates
boehmes
parents: 46736
diff changeset
   243
    |-> invoke name command
8e365bc843e9 fine-tune intended failure of smt method when the chosen SMT solver is not installed: restrict failures to true invocations of the SMT solver and don't fail for runs based on certificates
boehmes
parents: 46736
diff changeset
   244
    |> reconstruct ctxt
8e365bc843e9 fine-tune intended failure of smt method when the chosen SMT solver is not installed: restrict failures to true invocations of the SMT solver and don't fail for runs based on certificates
boehmes
parents: 46736
diff changeset
   245
    |>> distinct (op =)
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   246
  end
40166
d3bc972b7d9d optionally force the remote version of an SMT solver to be executed
boehmes
parents: 40165
diff changeset
   247
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   248
fun apply_solver ctxt = gen_apply o gen_preprocess 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
   249
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
   250
val default_max_relevant = #default_max_relevant oo get_info
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   251
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   252
val supports_filter = #supports_filter o snd o name_and_info_of 
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   253
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   254
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   255
(* check well-sortedness *)
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
   256
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
   257
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
   258
    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
   259
  | 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
   260
  | _ => 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
   261
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   262
(* 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
   263
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
   264
  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
   265
    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
   266
      "contains the universal sort {}"))
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   267
  else ()
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   268
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   269
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   270
(* filter *)
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
   271
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
   272
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
   273
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   274
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
   275
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   276
type 'a smt_filter_data = ('a * thm) list * ((int * thm) list * Proof.context)
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   277
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   278
fun smt_filter_preprocess st xwthms 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
   279
  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
   280
    val ctxt =
40199
2963511e121c use proper context
boehmes
parents: 40198
diff changeset
   281
      Proof.context_of st
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   282
      |> Config.put SMT_Config.oracle false
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   283
      |> Config.put SMT_Config.drop_bad_facts true
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41300
diff changeset
   284
      |> Config.put SMT_Config.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
   285
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
   286
    val {facts, goal, ...} = Proof.goal st
40357
82ebdd19c4a4 simulate more closely the behaviour of the tactic
boehmes
parents: 40332
diff changeset
   287
    val ({context=ctxt', prems, concl, ...}, _) = Subgoal.focus ctxt i goal
46497
89ccf66aa73d renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents: 46464
diff changeset
   288
    fun negate ct = Thm.dest_comb ct ||> Thm.apply cnot |-> Thm.apply
42320
1f09e4c680fc check if negating the goal is possible (avoids CTERM exception for Pure propositions)
boehmes
parents: 41761
diff changeset
   289
    val cprop =
1f09e4c680fc check if negating the goal is possible (avoids CTERM exception for Pure propositions)
boehmes
parents: 41761
diff changeset
   290
      (case try negate (Thm.rhs_of (SMT_Normalize.atomize_conv ctxt' concl)) of
1f09e4c680fc check if negating the goal is possible (avoids CTERM exception for Pure propositions)
boehmes
parents: 41761
diff changeset
   291
        SOME ct => ct
1f09e4c680fc check if negating the goal is possible (avoids CTERM exception for Pure propositions)
boehmes
parents: 41761
diff changeset
   292
      | NONE => raise SMT_Failure.SMT (SMT_Failure.Other_Failure (
1f09e4c680fc check if negating the goal is possible (avoids CTERM exception for Pure propositions)
boehmes
parents: 41761
diff changeset
   293
          "goal is not a HOL term")))
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
   294
  in
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   295
    map snd xwthms
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   296
    |> map_index I
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   297
    |> append (map (pair ~1 o pair NONE) (Thm.assume cprop :: prems @ facts))
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   298
    |> tap check_topsort
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   299
    |> gen_preprocess ctxt'
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   300
    |> pair (map (apsnd snd) xwthms)
41239
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   301
  end
d6e804ff29c3 split "smt_filter" into head and tail
blanchet
parents: 41131
diff changeset
   302
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   303
fun smt_filter_apply time_limit (xthms, (ithms, ctxt)) =
41241
7d07736aaaf6 make timeout part of the SMT filter's tail
blanchet
parents: 41239
diff changeset
   304
  let
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   305
    val ctxt' =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   306
      ctxt
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   307
      |> Config.put SMT_Config.timeout (Time.toReal time_limit)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   308
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   309
    fun filter_thms false = K xthms
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   310
      | filter_thms true = map_filter (try (nth xthms)) o fst
41241
7d07736aaaf6 make timeout part of the SMT filter's tail
blanchet
parents: 41239
diff changeset
   311
  in
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   312
    (ithms, ctxt')
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   313
    |> gen_apply
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   314
    |> filter_thms (supports_filter 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
   315
    |> 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
   316
  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
   317
  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
   318
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
   319
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   320
(* SMT tactic *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   321
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   322
local
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   323
  fun trace_assumptions ctxt iwthms idxs =
40165
boehmes
parents: 40164
diff changeset
   324
    let
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   325
      val wthms =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   326
        idxs
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   327
        |> filter (fn i => i >= 0)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   328
        |> map_filter (AList.lookup (op =) iwthms)
40165
boehmes
parents: 40164
diff changeset
   329
    in
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   330
      if Config.get ctxt SMT_Config.trace_used_facts andalso length wthms > 0
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   331
      then
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   332
        tracing (Pretty.string_of (Pretty.big_list "SMT used facts:"
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   333
          (map (Display.pretty_thm ctxt o snd) wthms)))
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   334
      else ()
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   335
    end
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   336
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   337
  fun solve ctxt iwthms =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   338
    iwthms
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   339
    |> tap check_topsort
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   340
    |> apply_solver ctxt
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   341
    |>> trace_assumptions ctxt iwthms
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   342
    |> snd
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   343
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   344
  fun str_of ctxt fail =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   345
    SMT_Failure.string_of_failure ctxt fail
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   346
    |> prefix ("Solver " ^ SMT_Config.solver_of ctxt ^ ": ")
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   347
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   348
  fun safe_solve ctxt iwthms = SOME (solve ctxt iwthms)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   349
    handle
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   350
      SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) =>
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   351
        (SMT_Config.verbose_msg ctxt (str_of ctxt) fail; NONE)
41761
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
   352
    | SMT_Failure.SMT (fail as SMT_Failure.Time_Out) =>
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
   353
        error ("SMT: Solver " ^ quote (SMT_Config.solver_of ctxt) ^ ": " ^
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
   354
          SMT_Failure.string_of_failure ctxt fail ^ " (setting the " ^
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42320
diff changeset
   355
          "configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)")
41761
2dc75bae5226 more explicit errors to inform users about problems related to SMT solvers;
boehmes
parents: 41601
diff changeset
   356
    | SMT_Failure.SMT fail => error (str_of ctxt fail)
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   357
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   358
  fun tag_rules thms = map_index (apsnd (pair NONE)) thms
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   359
  fun tag_prems thms = map (pair ~1 o pair NONE) thms
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   360
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   361
  fun resolve (SOME thm) = Tactic.rtac thm 1
46464
4cf5a84e2c05 normalized aliases;
wenzelm
parents: 42616
diff changeset
   362
    | resolve NONE = no_tac
41432
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   363
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   364
  fun tac prove ctxt rules =
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   365
    CONVERSION (SMT_Normalize.atomize_conv ctxt)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   366
    THEN' Tactic.rtac @{thm ccontr}
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   367
    THEN' SUBPROOF (fn {context, prems, ...} =>
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   368
      resolve (prove context (tag_rules rules @ tag_prems prems))) ctxt
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   369
in
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   370
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   371
val smt_tac = tac safe_solve
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   372
val smt_tac' = tac (SOME oo solve)
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   373
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   374
end
3214c39777ab differentiate between local and remote SMT solvers (e.g., "z3" vs. "remote_z3");
boehmes
parents: 41328
diff changeset
   375
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
   376
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   377
val smt_method =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   378
  Scan.optional Attrib.thms [] >>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   379
  (fn thms => fn ctxt => METHOD (fn facts =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   380
    HEADGOAL (smt_tac ctxt (thms @ facts))))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   381
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   382
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   383
(* setup *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   384
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   385
val setup =
38808
89ae86205739 more antiquotations;
wenzelm
parents: 36960
diff changeset
   386
  Method.setup @{binding smt} smt_method
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   387
    "Applies an SMT solver to the current goal."
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   388
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   389
end