src/HOL/Tools/SMT/smt_solver.ML
author wenzelm
Wed, 23 Jul 2025 14:53:21 +0200
changeset 82898 89da4dcd1fa8
parent 82092 93195437fdee
permissions -rw-r--r--
clarified colors, following d6a14ed060fb;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
     1
(*  Title:      HOL/Tools/SMT/smt_solver.ML
56078
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
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
     7
signature SMT_SOLVER =
56078
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*)
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 70293
diff changeset
    10
  datatype outcome = Unsat | Sat | Unknown | Time_Out
57163
blanchet
parents: 57159
diff changeset
    11
blanchet
parents: 57159
diff changeset
    12
  type parsed_proof =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    13
    {outcome: SMT_Failure.failure option,
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59632
diff changeset
    14
     fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list option,
57163
blanchet
parents: 57159
diff changeset
    15
     atp_proof: unit -> (term, string) ATP_Proof.atp_step list}
blanchet
parents: 57159
diff changeset
    16
blanchet
parents: 57159
diff changeset
    17
  type solver_config =
blanchet
parents: 57159
diff changeset
    18
    {name: string,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    19
     class: Proof.context -> SMT_Util.class,
57163
blanchet
parents: 57159
diff changeset
    20
     avail: unit -> bool,
blanchet
parents: 57159
diff changeset
    21
     command: unit -> string list,
blanchet
parents: 57159
diff changeset
    22
     options: Proof.context -> string list,
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57229
diff changeset
    23
     smt_options: (string * string) list,
77269
bc43f86c9598 added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents: 75339
diff changeset
    24
     good_slices: ((int * bool * bool * int * string) * string list) list,
57163
blanchet
parents: 57159
diff changeset
    25
     outcome: string -> string list -> outcome * string list,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    26
     parse_proof: (Proof.context -> SMT_Translate.replay_data ->
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
    27
       ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
    28
       parsed_proof) option,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    29
     replay: (Proof.context -> SMT_Translate.replay_data -> string list -> thm) option}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    30
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
  (*registry*)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    32
  val add_solver: solver_config -> theory -> theory
77269
bc43f86c9598 added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents: 75339
diff changeset
    33
  val good_slices: Proof.context -> string ->
bc43f86c9598 added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents: 75339
diff changeset
    34
    ((int * bool * bool * int * string) * string list) list
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    36
  (*filter*)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    37
  val smt_filter: Proof.context -> thm -> ((string * ATP_Problem_Generate.stature) * thm) list ->
81610
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
    38
    int -> Time.time -> ((string -> string) -> string -> string) -> string list -> parsed_proof
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    39
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    40
  (*tactic*)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    41
  val smt_tac: Proof.context -> thm list -> int -> tactic
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    42
  val smt_tac': Proof.context -> thm list -> int -> tactic
78461
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    43
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    44
  (*solver information*)
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    45
  type solver_info = {
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    46
    command: unit -> string list,
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    47
    smt_options: (string * string) list,
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    48
    good_slices: ((int * bool * bool * int * string) * string list) list,
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    49
    parse_proof: Proof.context -> SMT_Translate.replay_data ->
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    50
      ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    51
      parsed_proof,
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    52
    replay: Proof.context -> SMT_Translate.replay_data -> string list -> thm}
71713dd09c35 support for let in Alethe name bindings;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 77269
diff changeset
    53
  val name_and_info_of: Proof.context -> string * solver_info
57229
blanchet
parents: 57165
diff changeset
    54
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    55
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    56
structure SMT_Solver: SMT_SOLVER =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    58
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    59
(* interface to external solvers *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    60
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    61
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    62
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    63
fun with_trace ctxt msg f x =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    64
  let val _ = SMT_Config.trace_msg ctxt (fn () => msg) ()
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    65
  in f x end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
82090
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
    67
fun run ctxt name cmd input =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    68
  (case SMT_Config.certificates_of ctxt of
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    69
    NONE =>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    70
      if not (SMT_Config.is_available ctxt name) then
56131
blanchet
parents: 56128
diff changeset
    71
        error ("The SMT solver " ^ quote name ^ " is not installed")
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    72
      else if Config.get ctxt SMT_Config.debug_files = "" then
82090
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
    73
        with_trace ctxt ("Invoking SMT solver " ^ quote name ^ " ...") (Cache_IO.run cmd) input
56078
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
        let
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    76
          val base_path = Path.explode (Config.get ctxt SMT_Config.debug_files)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    77
          val in_path = Path.ext "smt_in" base_path
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    78
          val out_path = Path.ext "smt_out" base_path
82090
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
    79
          val _ = File.write in_path input
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
    80
          val result = Cache_IO.run cmd input
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
    81
          val _ = Bytes.write out_path (Bytes.terminate_lines (Process_Result.out_lines result))
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
    82
        in result end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    83
  | SOME certs =>
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    84
      (case Cache_IO.lookup certs input of
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    85
        (NONE, key) =>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
    86
          if Config.get ctxt SMT_Config.read_only_certificates then
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    87
            error ("Bad certificate cache: missing certificate")
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
          else
82090
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
    89
            Cache_IO.run_and_cache certs key cmd input
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    90
      | (SOME output, _) =>
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
    91
          with_trace ctxt ("Using cached certificate from " ^
62549
9498623b27f0 File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
wenzelm
parents: 61268
diff changeset
    92
            Path.print (Cache_IO.cache_path_of certs) ^ " ...") I output))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    93
74691
634e2323b6cf proper support of verit's return code for timeout
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    94
(* Z3 and cvc returns 1 if "get-proof" or "get-model" fails.
634e2323b6cf proper support of verit's return code for timeout
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    95
veriT returns 255 in that case and 14 for timeouts. *)
634e2323b6cf proper support of verit's return code for timeout
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    96
fun normal_return_codes "z3" = [0, 1]
634e2323b6cf proper support of verit's return code for timeout
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    97
  | normal_return_codes "verit" = [0, 14, 255]
634e2323b6cf proper support of verit's return code for timeout
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    98
  | normal_return_codes _ = [0, 1]
56087
2cd8fcb4804d slacker error code policy for Z3
blanchet
parents: 56086
diff changeset
    99
82090
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
   100
fun run_solver ctxt name cmd input =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   101
  let
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   102
    fun pretty tag lines = Pretty.string_of (Pretty.big_list tag (map Pretty.str lines))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   103
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   104
    val _ = SMT_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
82087
aee15b076916 clarified signature: re-use existing Process_Result.T (NB: err_lines are de-facto program startup errors, anything else is redirected to out_lines);
wenzelm
parents: 81610
diff changeset
   106
    val ({elapsed, ...}, result) =
82090
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
   107
      Timing.timing (SMT_Config.with_timeout ctxt (run ctxt name cmd)) input
82087
aee15b076916 clarified signature: re-use existing Process_Result.T (NB: err_lines are de-facto program startup errors, anything else is redirected to out_lines);
wenzelm
parents: 81610
diff changeset
   108
    val res = Process_Result.out_lines result
aee15b076916 clarified signature: re-use existing Process_Result.T (NB: err_lines are de-facto program startup errors, anything else is redirected to out_lines);
wenzelm
parents: 81610
diff changeset
   109
    val err = Process_Result.err_lines result
aee15b076916 clarified signature: re-use existing Process_Result.T (NB: err_lines are de-facto program startup errors, anything else is redirected to out_lines);
wenzelm
parents: 81610
diff changeset
   110
    val return_code = Process_Result.rc result
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   111
    val _ = SMT_Config.trace_msg ctxt (pretty "Solver:") err
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   112
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 66661
diff changeset
   113
    val output = drop_suffix (equal "") res
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   114
    val _ = SMT_Config.trace_msg ctxt (pretty "Result:") output
82092
93195437fdee clarified signature;
wenzelm
parents: 82091
diff changeset
   115
    val _ = SMT_Config.trace_msg ctxt (pretty "Time:") [Time.message elapsed]
93195437fdee clarified signature;
wenzelm
parents: 82091
diff changeset
   116
    val _ = SMT_Config.statistics_msg ctxt (pretty "Time:") [Time.message elapsed]
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   117
74691
634e2323b6cf proper support of verit's return code for timeout
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
   118
    val _ = member (op =) (normal_return_codes name) return_code orelse
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   119
      raise SMT_Failure.SMT (SMT_Failure.Abnormal_Termination return_code)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   120
  in output end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   121
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   122
fun trace_assms ctxt =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   123
  SMT_Config.trace_msg ctxt (Pretty.string_of o
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61033
diff changeset
   124
    Pretty.big_list "Assertions:" o map (Thm.pretty_thm ctxt o snd))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   125
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   126
fun trace_replay_data ({context = ctxt, typs, terms, ...} : SMT_Translate.replay_data) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   127
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   128
    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
   129
    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
   130
    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
   131
  in
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   132
    SMT_Config.trace_msg ctxt (fn () =>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   133
      Pretty.string_of (Pretty.big_list "Names:" [
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   134
        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
   135
        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
   136
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   137
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   138
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   139
82091
wenzelm
parents: 82090
diff changeset
   140
fun invoke memoize_fun_call name command cmd_options smt_options ithms ctxt =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   141
  let
82091
wenzelm
parents: 82090
diff changeset
   142
    val options = cmd_options @ SMT_Config.solver_options_of ctxt
80910
406a85a25189 clarified signature: more explicit operations;
wenzelm
parents: 78461
diff changeset
   143
    val comments = [implode_space options]
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   144
82091
wenzelm
parents: 82090
diff changeset
   145
    val (input, replay_data as {context = ctxt', ...}) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   146
      ithms
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   147
      |> tap (trace_assms ctxt)
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74691
diff changeset
   148
      |> SMT_Translate.translate ctxt name smt_options comments
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   149
      ||> tap trace_replay_data
81610
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   150
82090
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
   151
    val cmd = Bash.script (Bash.strings (command () @ options))
023845c29d48 avoid tmp files to improve robustness on Windows, where ERROR_PATH_NOT_FOUND has been seen (see also 569135d7352a);
wenzelm
parents: 82087
diff changeset
   152
    val run_cmd = run_solver ctxt' name cmd
81610
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   153
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   154
    val output_lines =
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   155
      (case memoize_fun_call of
82091
wenzelm
parents: 82090
diff changeset
   156
        NONE => run_cmd input
wenzelm
parents: 82090
diff changeset
   157
      | SOME memoize => split_lines (memoize (cat_lines o run_cmd) input))
81610
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   158
  in (output_lines, replay_data) end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   159
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   160
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   161
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   162
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   163
(* configuration *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   164
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 70293
diff changeset
   165
datatype outcome = Unsat | Sat | Unknown | Time_Out
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
57163
blanchet
parents: 57159
diff changeset
   167
type parsed_proof =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   168
  {outcome: SMT_Failure.failure option,
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59632
diff changeset
   169
   fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list option,
57163
blanchet
parents: 57159
diff changeset
   170
   atp_proof: unit -> (term, string) ATP_Proof.atp_step list}
blanchet
parents: 57159
diff changeset
   171
blanchet
parents: 57159
diff changeset
   172
type solver_config =
blanchet
parents: 57159
diff changeset
   173
  {name: string,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   174
   class: Proof.context -> SMT_Util.class,
57163
blanchet
parents: 57159
diff changeset
   175
   avail: unit -> bool,
blanchet
parents: 57159
diff changeset
   176
   command: unit -> string list,
blanchet
parents: 57159
diff changeset
   177
   options: Proof.context -> string list,
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57229
diff changeset
   178
   smt_options: (string * string) list,
77269
bc43f86c9598 added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents: 75339
diff changeset
   179
   good_slices: ((int * bool * bool * int * string) * string list) list,
57163
blanchet
parents: 57159
diff changeset
   180
   outcome: string -> string list -> outcome * string list,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   181
   parse_proof: (Proof.context -> SMT_Translate.replay_data ->
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   182
     ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   183
     parsed_proof) option,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   184
   replay: (Proof.context -> SMT_Translate.replay_data -> string list -> thm) option}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   185
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   186
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   187
(* check well-sortedness *)
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   188
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   189
val has_topsort = Term.exists_type (Term.exists_subtype (fn
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   190
    TFree (_, []) => true
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   191
  | TVar (_, []) => true
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   192
  | _ => false))
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   193
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   194
(* top sorts cause problems with atomization *)
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   195
fun check_topsort ctxt thm =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   196
  if has_topsort (Thm.prop_of thm) then (SMT_Normalize.drop_fact_warning ctxt thm; TrueI) else thm
56106
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   197
9cfea3ab002a simplify index handling
blanchet
parents: 56104
diff changeset
   198
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   199
(* registry *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   200
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   201
type solver_info = {
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   202
  command: unit -> string list,
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57229
diff changeset
   203
  smt_options: (string * string) list,
77269
bc43f86c9598 added refute mode to Sledgehammer to find 'counterexamples'
blanchet
parents: 75339
diff changeset
   204
  good_slices: ((int * bool * bool * int * string) * string list) list,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   205
  parse_proof: Proof.context -> SMT_Translate.replay_data ->
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   206
    ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   207
    parsed_proof,
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   208
  replay: Proof.context -> SMT_Translate.replay_data -> string list -> thm}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   209
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   210
structure Solvers = Generic_Data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   211
(
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   212
  type T = solver_info Symtab.table
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   213
  val empty = Symtab.empty
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   214
  fun merge data = Symtab.merge (K true) data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   215
)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   216
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   217
local
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   218
  fun parse_proof outcome parse_proof0 outer_ctxt replay_data xfacts prems concl output =
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   219
    (case outcome output of
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   220
      (Unsat, lines) =>
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   221
        (case parse_proof0 of
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   222
          SOME pp => pp outer_ctxt replay_data xfacts prems concl lines
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59632
diff changeset
   223
        | NONE => {outcome = NONE, fact_ids = NONE, atp_proof = K []})
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 70293
diff changeset
   224
    | (Time_Out, _) => raise SMT_Failure.SMT (SMT_Failure.Time_Out)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   225
    | (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat)))
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   226
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   227
  fun replay outcome replay0 oracle outer_ctxt
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   228
      (replay_data as {context = ctxt, ...} : SMT_Translate.replay_data) output =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   229
    (case outcome output of
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   230
      (Unsat, lines) =>
66661
fdab65297bd6 real oracle
blanchet
parents: 64304
diff changeset
   231
        if Config.get ctxt SMT_Config.oracle then
fdab65297bd6 real oracle
blanchet
parents: 64304
diff changeset
   232
          oracle ()
fdab65297bd6 real oracle
blanchet
parents: 64304
diff changeset
   233
        else
fdab65297bd6 real oracle
blanchet
parents: 64304
diff changeset
   234
          (case replay0 of
fdab65297bd6 real oracle
blanchet
parents: 64304
diff changeset
   235
            SOME replay => replay outer_ctxt replay_data lines
fdab65297bd6 real oracle
blanchet
parents: 64304
diff changeset
   236
          | NONE => error "No proof reconstruction for solver -- \
fdab65297bd6 real oracle
blanchet
parents: 64304
diff changeset
   237
            \declare [[smt_oracle]] to allow oracle")
70327
c04d4951a155 handle timeouts gracefully in 'smt' proof method (patch due to Mathias Fleury)
blanchet
parents: 70293
diff changeset
   238
    | (Time_Out, _) => raise SMT_Failure.SMT (SMT_Failure.Time_Out)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   239
    | (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat)))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   240
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69205
diff changeset
   241
  val cfalse = Thm.cterm_of \<^context> \<^prop>\<open>False\<close>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   242
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   243
75029
dc6769b86fd6 crude implementation of centralized slicing
blanchet
parents: 74817
diff changeset
   244
fun add_solver ({name, class, avail, command, options, smt_options, good_slices, outcome,
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   245
    parse_proof = parse_proof0, replay = replay0} : solver_config) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   246
  let
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   247
    fun solver oracle = {
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   248
      command = command,
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57229
diff changeset
   249
      smt_options = smt_options,
75029
dc6769b86fd6 crude implementation of centralized slicing
blanchet
parents: 74817
diff changeset
   250
      good_slices = good_slices,
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   251
      parse_proof = parse_proof (outcome name) parse_proof0,
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   252
      replay = replay (outcome name) replay0 oracle}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   253
56132
64eeda68e693 delayed construction of command (and of noncommercial check) + tuning
blanchet
parents: 56131
diff changeset
   254
    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
   255
  in
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   256
    Thm.add_oracle (Binding.name name, K cfalse) #-> (fn (_, oracle) =>
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   257
    Context.theory_map (Solvers.map (Symtab.update_new (name, solver oracle)))) #>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   258
    Context.theory_map (SMT_Config.add_solver info)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   259
  end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   260
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   261
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   262
56125
blanchet
parents: 56112
diff changeset
   263
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
   264
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   265
fun name_and_info_of ctxt =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   266
  let val name = SMT_Config.solver_of ctxt
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   267
  in (name, get_info ctxt name) end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   268
75029
dc6769b86fd6 crude implementation of centralized slicing
blanchet
parents: 74817
diff changeset
   269
val good_slices = #good_slices oo get_info
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   270
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57164
diff changeset
   271
fun apply_solver_and_replay ctxt thms0 =
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   272
  let
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 75063
diff changeset
   273
    val thms = map (pair SMT_Util.Axiom o check_topsort ctxt) thms0
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57229
diff changeset
   274
    val (name, {command, smt_options, replay, ...}) = name_and_info_of ctxt
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57229
diff changeset
   275
    val (output, replay_data) =
81610
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   276
      invoke NONE name command [] smt_options (SMT_Normalize.normalize ctxt thms) ctxt
57157
87b4d54b1fbe split replay and proof parsing for Z3
blanchet
parents: 57156
diff changeset
   277
  in replay ctxt replay_data output end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   278
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   279
72355
1f959abe99d5 Add more tacing to sledgehammer_isar_trace
desharna
parents: 72278
diff changeset
   280
(* filter (for Sledgehammer) *)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   281
81610
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   282
fun smt_filter ctxt0 goal xfacts i time_limit memoize_fun_call options =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   283
  let
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   284
    val ctxt = ctxt0 |> Config.put SMT_Config.timeout (Time.toReal time_limit)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   285
60695
757549b4bbe6 Variable.focus etc.: optional bindings provided by user;
wenzelm
parents: 60201
diff changeset
   286
    val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i NONE goal
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69205
diff changeset
   287
    fun negate ct = Thm.dest_comb ct ||> Thm.apply \<^cterm>\<open>Not\<close> |-> Thm.apply
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   288
    val cprop =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   289
      (case try negate (Thm.rhs_of (SMT_Normalize.atomize_conv ctxt concl)) of
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   290
        SOME ct => ct
61033
fd7fe96ca7b9 generate proper error instead of exception if goal cannot be atomized
blanchet
parents: 60752
diff changeset
   291
      | NONE => raise SMT_Failure.SMT (SMT_Failure.Other_Failure "cannot atomize goal"))
56082
ffd99d397a9f tuned ML interface
blanchet
parents: 56078
diff changeset
   292
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57164
diff changeset
   293
    val conjecture = Thm.assume cprop
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 75063
diff changeset
   294
    val thms =
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 75063
diff changeset
   295
      (SMT_Util.Conjecture, conjecture) ::
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 75063
diff changeset
   296
      map (pair SMT_Util.Hypothesis) prems @
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 75063
diff changeset
   297
      map (pair SMT_Util.Axiom o snd) xfacts
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 75063
diff changeset
   298
      |> map (apsnd (check_topsort ctxt))
57159
24cbdebba35a refactored Z3 to Isar proof construction code
blanchet
parents: 57158
diff changeset
   299
57239
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57229
diff changeset
   300
    val (name, {command, smt_options, parse_proof, ...}) = name_and_info_of ctxt
a40edeaa01b1 don't ask proof-disabled solvers to do proofs
blanchet
parents: 57229
diff changeset
   301
    val (output, replay_data) =
81610
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   302
      invoke (SOME memoize_fun_call) name command options smt_options
ed9ffd8e9e40 added option "cache_dir" to Sledgehammer
desharna
parents: 80910
diff changeset
   303
        (SMT_Normalize.normalize ctxt thms) ctxt
57164
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   304
  in
eb5f27ec3987 make SMT code less dependent on Z3 proofs
blanchet
parents: 57163
diff changeset
   305
    parse_proof ctxt replay_data xfacts (map Thm.prop_of prems) (Thm.term_of concl) output
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   306
  end
60201
90e88e521e0e made CVC4 support work also without unsat cores
blanchet
parents: 59632
diff changeset
   307
  handle SMT_Failure.SMT fail => {outcome = SOME fail, fact_ids = NONE, atp_proof = K []}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   308
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   309
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   310
(* SMT tactic *)
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
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   313
  fun str_of ctxt fail =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   314
    "Solver " ^ SMT_Config.solver_of ctxt ^ ": " ^ SMT_Failure.string_of_failure fail
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   315
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 57164
diff changeset
   316
  fun safe_solve ctxt facts = SOME (apply_solver_and_replay ctxt facts)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   317
    handle
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   318
      SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) =>
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   319
        (SMT_Config.verbose_msg ctxt (str_of ctxt) fail; NONE)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   320
    | SMT_Failure.SMT (fail as SMT_Failure.Time_Out) =>
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72355
diff changeset
   321
        (SMT_Config.verbose_msg ctxt (K ("SMT: Solver " ^ quote (SMT_Config.solver_of ctxt) ^ ": " ^
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   322
          SMT_Failure.string_of_failure fail ^ " (setting the " ^
72458
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72355
diff changeset
   323
          "configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)")) ();
b44e894796d5 add reconstruction for the SMT solver veriT
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 72355
diff changeset
   324
         NONE)
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   325
    | SMT_Failure.SMT fail => error (str_of ctxt fail)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   326
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60696
diff changeset
   327
  fun resolve ctxt (SOME thm) = resolve_tac ctxt [thm] 1
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60696
diff changeset
   328
    | resolve _ NONE = no_tac
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   329
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   330
  fun tac prove ctxt rules =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   331
    CONVERSION (SMT_Normalize.atomize_conv ctxt)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60696
diff changeset
   332
    THEN' resolve_tac ctxt @{thms ccontr}
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60696
diff changeset
   333
    THEN' SUBPROOF (fn {context = ctxt', prems, ...} =>
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60696
diff changeset
   334
      resolve ctxt' (prove ctxt' (rules @ prems))) ctxt
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   335
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   336
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   337
val smt_tac = tac safe_solve
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57239
diff changeset
   338
val smt_tac' = tac (SOME oo apply_solver_and_replay)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   339
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   340
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   341
57229
blanchet
parents: 57165
diff changeset
   342
end;