src/HOL/Tools/SMT2/smt2_solver.ML
author blanchet
Mon, 02 Jun 2014 17:34:26 +0200
changeset 57157 87b4d54b1fbe
parent 57156 3546a67226ea
child 57158 f028d93798e6
permissions -rw-r--r--
split replay and proof parsing for Z3
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/SMT2/smt2_solver.ML
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     3
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     4
SMT solvers registry and SMT tactic.
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     5
*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     6
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     7
signature SMT2_SOLVER =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     8
sig
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     9
  (*configuration*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    10
  datatype outcome = Unsat | Sat | Unknown
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    11
  type solver_config = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    12
    name: string,
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56087
diff changeset
    13
    class: Proof.context -> SMT2_Util.class,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    14
    avail: unit -> bool,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    15
    command: unit -> string list,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    16
    options: Proof.context -> string list,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
    default_max_relevant: int,
56125
blanchet
parents: 56112
diff changeset
    18
    can_filter: bool,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    19
    outcome: string -> string list -> outcome * string list,
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
    20
    parse_proof: (Proof.context -> SMT2_Translate.replay_data -> string list ->
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
    21
      (int * (int * thm)) list * Z3_New_Proof.z3_step list) option,
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
    22
    replay: (Proof.context -> SMT2_Translate.replay_data -> string list -> thm) option}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    23
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    24
  (*registry*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    25
  val add_solver: solver_config -> theory -> theory
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    26
  val default_max_relevant: Proof.context -> string -> int
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    27
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    28
  (*filter*)
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
    29
  val smt2_filter: Proof.context -> thm -> ('a * (int option * thm)) list -> int -> Time.time ->
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56125
diff changeset
    30
    {outcome: SMT2_Failure.failure option, rewrite_rules: thm list, conjecture_id: int,
56981
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
    31
     prem_ids: int list, helper_ids: (int * thm) list, fact_ids: (int * ('a * thm)) list,
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56125
diff changeset
    32
     z3_proof: Z3_New_Proof.z3_step list}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    33
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    34
  (*tactic*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
  val smt2_tac: Proof.context -> thm list -> int -> tactic
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    36
  val smt2_tac': Proof.context -> thm list -> int -> tactic
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    37
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    38
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    39
structure SMT2_Solver: SMT2_SOLVER =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    40
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    41
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    42
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    43
(* interface to external solvers *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    44
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    45
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    46
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    47
fun make_command command options problem_path proof_path =
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    48
  "(exec 2>&1;" :: map File.shell_quote (command () @ options) @
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    49
  [File.shell_path problem_path, ")", ">", File.shell_path proof_path]
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    50
  |> space_implode " "
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    51
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    52
fun with_trace ctxt msg f x =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    53
  let val _ = SMT2_Config.trace_msg ctxt (fn () => msg) ()
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    54
  in f x end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    55
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    56
fun run ctxt name mk_cmd input =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
  (case SMT2_Config.certificates_of ctxt of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    58
    NONE =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
      if not (SMT2_Config.is_available ctxt name) then
56131
blanchet
parents: 56128
diff changeset
    60
        error ("The SMT solver " ^ quote name ^ " is not installed")
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    61
      else if Config.get ctxt SMT2_Config.debug_files = "" then
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    62
        with_trace ctxt ("Invoking SMT solver " ^ quote name ^ " ...") (Cache_IO.run mk_cmd) input
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    63
      else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    64
        let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    65
          val base_path = Path.explode (Config.get ctxt SMT2_Config.debug_files)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
          val in_path = Path.ext "smt2_in" base_path
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    67
          val out_path = Path.ext "smt2_out" base_path
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    68
        in Cache_IO.raw_run mk_cmd input in_path out_path end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    69
  | SOME certs =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    70
      (case Cache_IO.lookup certs input of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    71
        (NONE, key) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    72
          if Config.get ctxt SMT2_Config.read_only_certificates then
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    73
            error ("Bad certificate cache: missing certificate")
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    74
          else
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    75
            Cache_IO.run_and_cache certs key mk_cmd input
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    76
      | (SOME output, _) =>
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    77
          with_trace ctxt ("Using cached certificate from " ^
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    78
            File.shell_path (Cache_IO.cache_path_of certs) ^ " ...") I output))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    79
56087
2cd8fcb4804d slacker error code policy for Z3
blanchet
parents: 56086
diff changeset
    80
(* Z3 returns 1 if "get-model" or "get-model" fails *)
2cd8fcb4804d slacker error code policy for Z3
blanchet
parents: 56086
diff changeset
    81
val normal_return_codes = [0, 1]
2cd8fcb4804d slacker error code policy for Z3
blanchet
parents: 56086
diff changeset
    82
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    83
fun run_solver ctxt name mk_cmd input =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    84
  let
56087
2cd8fcb4804d slacker error code policy for Z3
blanchet
parents: 56086
diff changeset
    85
    fun pretty tag ls = Pretty.string_of (Pretty.big_list tag (map Pretty.str ls))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    86
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    87
    val _ = SMT2_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    89
    val {redirected_output = res, output = err, return_code} =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    90
      SMT2_Config.with_timeout ctxt (run ctxt name mk_cmd) input
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    91
    val _ = SMT2_Config.trace_msg ctxt (pretty "Solver:") err
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    92
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    93
    val output = fst (take_suffix (equal "") res)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    94
    val _ = SMT2_Config.trace_msg ctxt (pretty "Result:") output
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    95
56087
2cd8fcb4804d slacker error code policy for Z3
blanchet
parents: 56086
diff changeset
    96
    val _ = member (op =) normal_return_codes return_code orelse
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    97
      raise SMT2_Failure.SMT (SMT2_Failure.Abnormal_Termination return_code)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    98
  in output end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    99
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   100
fun trace_assms ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   101
  SMT2_Config.trace_msg ctxt (Pretty.string_of o
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   102
    Pretty.big_list "Assertions:" o map (Display.pretty_thm ctxt o snd))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   103
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   104
fun trace_replay_data ({context = ctxt, typs, terms, ...} : SMT2_Translate.replay_data) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   106
    fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
    fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   108
    fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   109
  in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   110
    SMT2_Config.trace_msg ctxt (fn () =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   111
      Pretty.string_of (Pretty.big_list "Names:" [
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   112
        Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   113
        Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   114
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   115
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   116
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   117
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   118
fun invoke name command ithms ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   119
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   120
    val options = SMT2_Config.solver_options_of ctxt
56112
040424c3800d updated SMT2 examples and certificates
blanchet
parents: 56106
diff changeset
   121
    val comments = [space_implode " " options]
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   122
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   123
    val (str, replay_data as {context = ctxt', ...}) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   124
      ithms
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   125
      |> tap (trace_assms ctxt)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   126
      |> SMT2_Translate.translate ctxt comments
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   127
      ||> tap trace_replay_data
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   128
  in (run_solver ctxt' name (make_command command options) str, replay_data) end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   129
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   130
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   131
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   132
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   133
(* configuration *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   134
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   135
datatype outcome = Unsat | Sat | Unknown
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   136
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   137
type solver_config = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   138
  name: string,
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56087
diff changeset
   139
  class: Proof.context -> SMT2_Util.class,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   140
  avail: unit -> bool,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   141
  command: unit -> string list,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   142
  options: Proof.context -> string list,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   143
  default_max_relevant: int,
56125
blanchet
parents: 56112
diff changeset
   144
  can_filter: bool,
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   145
  outcome: string -> string list -> outcome * string list,
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   146
  parse_proof: (Proof.context -> SMT2_Translate.replay_data -> string list ->
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   147
    (int * (int * thm)) list * Z3_New_Proof.z3_step list) option,
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   148
  replay: (Proof.context -> SMT2_Translate.replay_data -> string list -> thm) option}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   149
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   150
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   151
(* check well-sortedness *)
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   152
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   153
val has_topsort = Term.exists_type (Term.exists_subtype (fn
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   154
    TFree (_, []) => true
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   155
  | TVar (_, []) => true
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   156
  | _ => false))
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   157
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   158
(* top sorts cause problems with atomization *)
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   159
fun check_topsort ctxt thm =
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   160
  if has_topsort (Thm.prop_of thm) then (SMT2_Normalize.drop_fact_warning ctxt thm; TrueI) else thm
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   161
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   162
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   163
(* registry *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   164
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
type solver_info = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
  command: unit -> string list,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   167
  default_max_relevant: int,
56125
blanchet
parents: 56112
diff changeset
   168
  can_filter: bool,
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   169
  parse_proof: Proof.context -> SMT2_Translate.replay_data -> string list ->
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   170
    (int * (int * thm)) list * Z3_New_Proof.z3_step list,
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   171
  replay: Proof.context -> SMT2_Translate.replay_data -> string list -> thm}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   172
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   173
structure Solvers = Generic_Data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   174
(
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   175
  type T = solver_info Symtab.table
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   176
  val empty = Symtab.empty
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   177
  val extend = I
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   178
  fun merge data = Symtab.merge (K true) data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   179
)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   180
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   181
local
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   182
  fun parse_proof outcome parse_proof0 outer_ctxt replay_data output =
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   183
    (case outcome output of
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   184
      (Unsat, ls) =>
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   185
        (case parse_proof0 of SOME pp => pp outer_ctxt replay_data ls | NONE => ([], []))
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   186
    | (result, _) =>
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   187
        raise SMT2_Failure.SMT (SMT2_Failure.Counterexample {
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   188
          is_real_cex = (result = Sat), free_constraints = [], const_defs = []}))
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   189
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   190
  fun replay outcome replay0 oracle outer_ctxt
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   191
      (replay_data as {context = ctxt, ...} : SMT2_Translate.replay_data) output =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   192
    (case outcome output of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   193
      (Unsat, ls) =>
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   194
        if not (Config.get ctxt SMT2_Config.oracle) andalso is_some replay0
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   195
        then the replay0 outer_ctxt replay_data ls
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   196
        else oracle ()
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   197
    | (result, _) =>
57156
3546a67226ea removed counterexample parser (obsolete and useless in practice)
blanchet
parents: 56981
diff changeset
   198
        raise SMT2_Failure.SMT (SMT2_Failure.Counterexample {
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   199
          is_real_cex = (result = Sat), free_constraints = [], const_defs = []}))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   200
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   201
  val cfalse = Thm.cterm_of @{theory} @{prop False}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   202
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   203
56125
blanchet
parents: 56112
diff changeset
   204
fun add_solver ({name, class, avail, command, options, default_max_relevant, can_filter,
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   205
    outcome, parse_proof = parse_proof0, replay = replay0} : solver_config) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   206
  let
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   207
    fun solver oracle = {
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   208
      command = command,
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   209
      default_max_relevant = default_max_relevant,
56125
blanchet
parents: 56112
diff changeset
   210
      can_filter = can_filter,
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   211
      parse_proof = parse_proof (outcome name) parse_proof0,
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   212
      replay = replay (outcome name) replay0 oracle}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   213
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   214
    val info = {name = name, class = class, avail = avail, options = options}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   215
  in
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   216
    Thm.add_oracle (Binding.name name, K cfalse) #-> (fn (_, oracle) =>
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   217
    Context.theory_map (Solvers.map (Symtab.update_new (name, solver oracle)))) #>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   218
    Context.theory_map (SMT2_Config.add_solver info)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   219
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   220
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   221
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   222
56125
blanchet
parents: 56112
diff changeset
   223
fun get_info ctxt name = the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   224
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   225
fun name_and_info_of ctxt =
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   226
  let val name = SMT2_Config.solver_of ctxt
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   227
  in (name, get_info ctxt name) end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   228
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   229
fun apply_solver_and_parse_proof ctxt wthms0 =
56082
ffd99d397a9f tuned ML interface
blanchet
parents: 56078
diff changeset
   230
  let
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   231
    val wthms = map (apsnd (check_topsort ctxt)) wthms0
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   232
    val (name, {command, parse_proof, ...}) = name_and_info_of ctxt
56125
blanchet
parents: 56112
diff changeset
   233
    val (output, replay_data) = invoke name command (SMT2_Normalize.normalize ctxt wthms) ctxt
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   234
  in (parse_proof ctxt replay_data output, replay_data) end
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   235
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   236
fun apply_solver_and_replay ctxt wthms0 =
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   237
  let
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   238
    val wthms = map (apsnd (check_topsort ctxt)) wthms0
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   239
    val (name, {command, replay, ...}) = name_and_info_of ctxt
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   240
    val (output, replay_data) = invoke name command (SMT2_Normalize.normalize ctxt wthms) ctxt
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   241
  in replay ctxt replay_data output end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   242
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   243
val default_max_relevant = #default_max_relevant oo get_info
56125
blanchet
parents: 56112
diff changeset
   244
val can_filter = #can_filter o snd o name_and_info_of 
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   245
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   246
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   247
(* filter *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   248
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   249
val no_id = ~1
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   250
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   251
fun smt2_filter ctxt goal xwfacts i time_limit =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   252
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   253
    val ctxt =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   254
      ctxt
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   255
      |> Config.put SMT2_Config.oracle false
56082
ffd99d397a9f tuned ML interface
blanchet
parents: 56078
diff changeset
   256
      |> Config.put SMT2_Config.timeout (Time.toReal time_limit)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   257
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   258
    val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i goal
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   259
    fun negate ct = Thm.dest_comb ct ||> Thm.apply @{cterm Not} |-> Thm.apply
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   260
    val cprop =
56082
ffd99d397a9f tuned ML interface
blanchet
parents: 56078
diff changeset
   261
      (case try negate (Thm.rhs_of (SMT2_Normalize.atomize_conv ctxt concl)) of
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   262
        SOME ct => ct
56094
2adbc6e4cd8f let exception pass through in debug mode
blanchet
parents: 56090
diff changeset
   263
      | NONE => raise SMT2_Failure.SMT (SMT2_Failure.Other_Failure "goal is not a HOL term"))
56082
ffd99d397a9f tuned ML interface
blanchet
parents: 56078
diff changeset
   264
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   265
    val wconjecture = (NONE, Thm.assume cprop)
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   266
    val wprems = map (pair NONE) prems
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   267
    val wfacts = map snd xwfacts
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   268
    val wthms = wconjecture :: wprems @ wfacts
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   269
    val iwthms = map_index I wthms
56082
ffd99d397a9f tuned ML interface
blanchet
parents: 56078
diff changeset
   270
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   271
    val conjecture_i = 0
56981
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
   272
    val prems_i = 1
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
   273
    val facts_i = prems_i + length wprems
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   274
  in
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   275
    wthms
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   276
    |> apply_solver_and_parse_proof ctxt
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   277
    |> (fn ((iidths0, z3_proof), {rewrite_rules, ...}) =>
56981
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
   278
      let
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
   279
        val iidths = if can_filter ctxt then iidths0 else map (apsnd (apfst (K no_id))) iwthms
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
   280
        fun id_of_index i = the_default no_id (Option.map fst (AList.lookup (op =) iidths i))
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   281
      in
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56125
diff changeset
   282
        {outcome = NONE,
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56125
diff changeset
   283
         rewrite_rules = rewrite_rules,
56981
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
   284
         conjecture_id = id_of_index conjecture_i,
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
   285
         prem_ids = map id_of_index (prems_i upto facts_i - 1),
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   286
         helper_ids = map_filter (try (fn (~1, idth) => idth)) iidths,
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   287
         fact_ids = map_filter (fn (i, (id, _)) =>
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   288
           try (apsnd (apsnd snd o nth xwfacts)) (id, i - facts_i)) iidths,
56104
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   289
         z3_proof = z3_proof}
fd6e132ee4fb correctly reconstruct helper facts (e.g. 'nat_int') in Isar proofs
blanchet
parents: 56099
diff changeset
   290
      end)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   291
  end
56128
c106ac2ff76d undo rewrite rules (e.g. for 'fun_app') in Isar
blanchet
parents: 56125
diff changeset
   292
  handle SMT2_Failure.SMT fail => {outcome = SOME fail, rewrite_rules = [], conjecture_id = no_id,
56981
3ef45ce002b5 honor original format of conjecture or hypotheses in Z3-to-Isar proofs
blanchet
parents: 56132
diff changeset
   293
    prem_ids = [], helper_ids = [], fact_ids = [], z3_proof = []}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   294
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   295
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   296
(* SMT tactic *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   297
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   298
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   299
  fun str_of ctxt fail =
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   300
    "Solver " ^ SMT2_Config.solver_of ctxt ^ ": " ^ SMT2_Failure.string_of_failure ctxt fail
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   301
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   302
  fun safe_solve ctxt wfacts = SOME (apply_solver_and_replay ctxt wfacts)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   303
    handle
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   304
      SMT2_Failure.SMT (fail as SMT2_Failure.Counterexample _) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   305
        (SMT2_Config.verbose_msg ctxt (str_of ctxt) fail; NONE)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   306
    | SMT2_Failure.SMT (fail as SMT2_Failure.Time_Out) =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   307
        error ("SMT: Solver " ^ quote (SMT2_Config.solver_of ctxt) ^ ": " ^
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   308
          SMT2_Failure.string_of_failure ctxt fail ^ " (setting the " ^
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   309
          "configuration option " ^ quote (Config.name_of SMT2_Config.timeout) ^ " might help)")
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   310
    | SMT2_Failure.SMT fail => error (str_of ctxt fail)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   311
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   312
  fun resolve (SOME thm) = rtac thm 1
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   313
    | resolve NONE = no_tac
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   314
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   315
  fun tac prove ctxt rules =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   316
    CONVERSION (SMT2_Normalize.atomize_conv ctxt)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   317
    THEN' rtac @{thm ccontr}
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   318
    THEN' SUBPROOF (fn {context = ctxt, prems, ...} =>
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   319
      resolve (prove ctxt (map (pair NONE) (rules @ prems)))) ctxt
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   320
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   321
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   322
val smt2_tac = tac safe_solve
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   323
val smt2_tac' = tac (SOME oo apply_solver_and_replay)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   324
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   325
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   326
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   327
end