src/HOL/Tools/SMT/smtlib_interface.ML
author Elisabeth Lempa
Fri, 19 Sep 2025 13:11:51 +0200
changeset 83191 76878779e355
parent 80910 406a85a25189
permissions -rw-r--r--
Added support for Vampire as an SMT solver. (experimental)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
     1
(*  Title:      HOL/Tools/SMT/smtlib_interface.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
    Author:     Jasmin Blanchette, TU Muenchen
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     4
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     5
Interface to SMT solvers based on the SMT-LIB 2 format.
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
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
     8
signature SMTLIB_INTERFACE =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
     9
sig
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    10
  val smtlibC: SMT_Util.class
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
    11
  val hosmtlibC: SMT_Util.class
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    12
  val bvsmlibC: SMT_Util.class
78177
ea7a3cc64df5 early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 75365
diff changeset
    13
  val realsmlibC: SMT_Util.class
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    14
  val add_logic: int * (string -> term list -> string option) -> Context.generic -> Context.generic
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    15
  val del_logic: int * (string -> term list -> string option) -> Context.generic -> Context.generic
83191
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
    16
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
    17
  type dtype_decls = (string * (string * (string * string) list) list) list
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
    18
  type func_decl = string * (string list * string)
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
    19
  val serialize: (BNF_Util.fp_kind * dtype_decls -> Buffer.T -> Buffer.T) -> (string * string) list -> string list ->
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
    20
    {dtyps: (BNF_Util.fp_kind * (string * (string * (string * string) list) list)) list,
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
    21
     funcs: func_decl list, logic: string, sorts: string list} ->
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
    22
    (SMT_Util.role * SMT_Translate.sterm) list -> string
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
    23
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
    24
  val translate_config: SMT_Util.order -> Proof.context -> SMT_Translate.config
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
    25
  val assert_name_of_role_and_index: SMT_Util.role -> int -> string
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
    26
  val role_and_index_of_assert_name: string -> SMT_Util.role * int
57229
blanchet
parents: 57203
diff changeset
    27
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    28
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    29
structure SMTLIB_Interface: SMTLIB_INTERFACE =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    30
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    31
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
    32
val hoC = ["ho"]
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
    33
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
    34
val smtlibC = ["smtlib"]   (* SMT-LIB 2 *)
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
    35
val hosmtlibC = smtlibC @ hoC   (* possibly SMT-LIB 3 *)
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    36
val bvsmlibC = smtlibC @ ["BV"] (* if BV are supported *)
78177
ea7a3cc64df5 early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 75365
diff changeset
    37
val realsmlibC = ["Real"]
56078
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
(* builtins *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    40
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    41
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    42
  fun int_num _ i = SOME (string_of_int i)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    43
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    44
  fun is_linear [t] = SMT_Util.is_number t
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    45
    | is_linear [t, u] = SMT_Util.is_number t orelse SMT_Util.is_number u
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    46
    | is_linear _ = false
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    47
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    48
  fun times _ _ ts =
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    49
    let val mk = Term.list_comb o pair \<^Const>\<open>times \<^Type>\<open>int\<close>\<close>
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    50
    in if is_linear ts then SOME ("*", 2, ts, mk) else NONE end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    51
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    53
val setup_builtins =
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
    54
  SMT_Builtin.add_builtin_typ hosmtlibC
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
    55
    (\<^typ>\<open>'a => 'b\<close>, fn Type (\<^type_name>\<open>fun\<close>, Ts) => SOME ("->", Ts), K (K NONE)) #>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    56
  fold (SMT_Builtin.add_builtin_typ smtlibC) [
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
    57
    (\<^typ>\<open>bool\<close>, K (SOME ("Bool", [])), K (K NONE)),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67405
diff changeset
    58
    (\<^typ>\<open>int\<close>, K (SOME ("Int", [])), int_num)] #>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    59
  fold (SMT_Builtin.add_builtin_fun' smtlibC) [
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    60
    (\<^Const>\<open>True\<close>, "true"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    61
    (\<^Const>\<open>False\<close>, "false"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    62
    (\<^Const>\<open>Not\<close>, "not"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    63
    (\<^Const>\<open>conj\<close>, "and"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    64
    (\<^Const>\<open>disj\<close>, "or"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    65
    (\<^Const>\<open>implies\<close>, "=>"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    66
    (\<^Const>\<open>HOL.eq \<^typ>\<open>'a\<close>\<close>, "="),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    67
    (\<^Const>\<open>If \<^typ>\<open>'a\<close>\<close>, "ite"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    68
    (\<^Const>\<open>less \<^Type>\<open>int\<close>\<close>, "<"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    69
    (\<^Const>\<open>less_eq \<^Type>\<open>int\<close>\<close>, "<="),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    70
    (\<^Const>\<open>uminus \<^Type>\<open>int\<close>\<close>, "-"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    71
    (\<^Const>\<open>plus \<^Type>\<open>int\<close>\<close>, "+"),
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    72
    (\<^Const>\<open>minus \<^Type>\<open>int\<close>\<close>, "-")] #>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    73
  SMT_Builtin.add_builtin_fun smtlibC
74382
8d0294d877bd clarified antiquotations;
wenzelm
parents: 69597
diff changeset
    74
    (Term.dest_Const \<^Const>\<open>times \<^Type>\<open>int\<close>\<close>, times)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    75
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    76
end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    77
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    78
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    79
(* serialization *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    80
57238
blanchet
parents: 57229
diff changeset
    81
(** logic **)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    82
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    83
fun fst_int_ord ((i1, _), (i2, _)) = int_ord (i1, i2)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    84
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    85
structure Logics = Generic_Data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    86
(
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    87
  type T = (int * (string -> term list -> string option)) list
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
  val empty = []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    89
  fun merge data = Ord_List.merge fst_int_ord data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    90
)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    91
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    92
fun add_logic pf = Logics.map (Ord_List.insert fst_int_ord pf)
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    93
fun del_logic pf = Logics.map (Ord_List.remove fst_int_ord pf)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    94
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    95
fun choose_logic ctxt prover ts =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    96
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    97
    fun choose [] = "AUFLIA"
74817
1fd8705503b4 generate problems with correct logic for veriT
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 74561
diff changeset
    98
      | choose ((_, f) :: fs) = (case f prover ts of SOME s => s | NONE => choose fs)
57553
2a6c31ac1c9a don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
blanchet
parents: 57239
diff changeset
    99
  in
2a6c31ac1c9a don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
blanchet
parents: 57239
diff changeset
   100
    (case choose (Logics.get (Context.Proof ctxt)) of
2a6c31ac1c9a don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
blanchet
parents: 57239
diff changeset
   101
      "" => "" (* for default Z3 logic, a subset of everything *)
2a6c31ac1c9a don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
blanchet
parents: 57239
diff changeset
   102
    | logic => "(set-logic " ^ logic ^ ")\n")
2a6c31ac1c9a don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
blanchet
parents: 57239
diff changeset
   103
  end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   104
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   105
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   106
(** serialization **)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   107
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   108
fun var i = "?v" ^ string_of_int i
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   109
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   110
fun tree_of_sterm l (SMT_Translate.SVar (i, [])) = SMTLIB.Sym (var (l - i - 1))
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   111
  | tree_of_sterm l (SMT_Translate.SVar (i, ts)) =
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   112
      SMTLIB.S (SMTLIB.Sym (var (l - i - 1)) :: map (tree_of_sterm l) ts)
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   113
  | tree_of_sterm _ (SMT_Translate.SConst (n, [])) = SMTLIB.Sym n
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   114
  | tree_of_sterm l (SMT_Translate.SConst (n, ts)) =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   115
      SMTLIB.S (SMTLIB.Sym n :: map (tree_of_sterm l) ts)
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   116
  | tree_of_sterm l (SMT_Translate.SQua (q, ss, pats, t)) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   117
      let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   118
        val l' = l + length ss
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   119
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   120
        fun quant_name SMT_Translate.SForall = "forall"
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   121
          | quant_name SMT_Translate.SExists = "exists"
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   122
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   123
        fun gen_trees_of_pat keyword ps =
59021
b29281d6d1db always generate patterns as lists of lists, to avoid confusing CVC4 (and stick to what SMT-LIB 2 actually says)
blanchet
parents: 59015
diff changeset
   124
          [SMTLIB.Key keyword, SMTLIB.S (map (tree_of_sterm l') ps)]
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   125
        fun trees_of_pat (SMT_Translate.SPat ps) = gen_trees_of_pat "pattern" ps
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   126
          | trees_of_pat (SMT_Translate.SNoPat ps) = gen_trees_of_pat "no-pattern" ps
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56091
diff changeset
   127
        fun tree_of_pats [] t = t
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   128
          | tree_of_pats pats t = SMTLIB.S (SMTLIB.Sym "!" :: t :: maps trees_of_pat pats)
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
        val vs = map_index (fn (i, ty) =>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   131
          SMTLIB.S [SMTLIB.Sym (var (l + i)), SMTLIB.Sym ty]) ss
56078
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
        val body = t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   134
          |> tree_of_sterm l'
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56091
diff changeset
   135
          |> tree_of_pats pats
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   136
      in
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   137
        SMTLIB.S [SMTLIB.Sym (quant_name q), SMTLIB.S vs, body]
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   138
      end
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   139
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   140
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   141
fun sctrarg (sel, typ) = "(" ^ sel ^ " " ^ typ ^ ")"
80910
406a85a25189 clarified signature: more explicit operations;
wenzelm
parents: 78177
diff changeset
   142
fun sctr (name, args) = enclose "(" ")" (implode_space (name :: map sctrarg args))
406a85a25189 clarified signature: more explicit operations;
wenzelm
parents: 78177
diff changeset
   143
fun sdatatype (name, ctrs) = enclose "(" ")" (implode_space (name :: map sctr ctrs))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   144
80910
406a85a25189 clarified signature: more explicit operations;
wenzelm
parents: 78177
diff changeset
   145
fun string_of_fun (f, (ss, s)) = f ^ " (" ^ implode_space ss ^ ") " ^ s
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   146
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   147
fun named_sterm s t = SMTLIB.S [SMTLIB.Sym "!", t, SMTLIB.Key "named", SMTLIB.Sym s]
57203
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
   148
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   149
val conjecture_prefix = "conjecture" (* FUDGE *)
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   150
val hypothesis_prefix = "hypothesis" (* FUDGE *)
78177
ea7a3cc64df5 early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 75365
diff changeset
   151
val axiom_prefix = "a" (* matching Alethe's convention *)
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   152
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   153
fun assert_prefix_of_role SMT_Util.Conjecture = conjecture_prefix
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   154
  | assert_prefix_of_role SMT_Util.Hypothesis = hypothesis_prefix
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   155
  | assert_prefix_of_role SMT_Util.Axiom = axiom_prefix
57203
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
   156
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   157
fun assert_name_of_role_and_index role i = assert_prefix_of_role role ^ string_of_int i
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   158
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   159
fun unprefix_axiom s =
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   160
  try (pair SMT_Util.Conjecture o unprefix conjecture_prefix) s
75365
83940294cc67 fixed generation of Isar proofs e89709b80b6e
desharna
parents: 75274
diff changeset
   161
  |> curry merge_options (try (pair SMT_Util.Hypothesis o unprefix hypothesis_prefix) s)
83940294cc67 fixed generation of Isar proofs e89709b80b6e
desharna
parents: 75274
diff changeset
   162
  |> curry merge_options (try (pair SMT_Util.Axiom o unprefix axiom_prefix) s)
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   163
  |> the
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   164
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   165
fun role_and_index_of_assert_name s =
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   166
  apsnd (the_default ~1 o Int.fromString) (unprefix_axiom s)
57203
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
   167
83191
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
   168
type dtype_decls = (string * (string * (string * string) list) list) list
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
   169
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
   170
fun sdtyp (fp, dtyps : dtype_decls) =
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   171
  Buffer.add (enclose ("(declare-" ^ BNF_FP_Util.co_prefix fp ^ "datatypes () (") "))\n"
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   172
    (space_implode "\n  " (map sdatatype dtyps)))
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   173
83191
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
   174
type func_decl = string * (string list * string)
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
   175
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
   176
fun serialize (sdtyp : BNF_Util.fp_kind * dtype_decls -> Buffer.T -> Buffer.T) smt_options comments
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
   177
      {logic, sorts, dtyps, funcs} ts =
59015
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   178
  let
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67399
diff changeset
   179
    val unsat_core = member (op =) smt_options (":produce-unsat-cores", "true")
59015
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   180
  in
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   181
    Buffer.empty
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   182
    |> fold (Buffer.add o enclose "; " "\n") comments
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   183
    |> fold (fn (k, v) => Buffer.add ("(set-option " ^ k ^ " " ^ v ^ ")\n")) smt_options
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   184
    |> Buffer.add logic
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   185
    |> fold (Buffer.add o enclose "(declare-sort " " 0)\n") (sort fast_string_ord sorts)
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   186
    |> fold sdtyp (AList.coalesce (op =) dtyps)
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   187
    |> fold (Buffer.add o enclose "(declare-fun " ")\n" o string_of_fun)
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 59021
diff changeset
   188
        (sort (fast_string_ord o apply2 fst) funcs)
75274
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   189
    |> fold (fn (i, (role, t)) => Buffer.add (enclose "(assert " ")\n"
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   190
        (SMTLIB.str_of (named_sterm (assert_name_of_role_and_index role i) (tree_of_sterm 0 t)))))
e89709b80b6e used more descriptive assert names in SMT-Lib output
desharna
parents: 74817
diff changeset
   191
        (map_index I ts)
59015
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   192
    |> Buffer.add "(check-sat)\n"
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   193
    |> Buffer.add (if unsat_core then "(get-unsat-core)\n" else "(get-proof)\n")
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   194
    |> Buffer.content
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   195
  end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   196
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   197
(* interface *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   198
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   199
fun translate_config order ctxt = {
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   200
  order = order,
57238
blanchet
parents: 57229
diff changeset
   201
  logic = choose_logic ctxt,
58360
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
   202
  fp_kinds = [],
83191
76878779e355 Added support for Vampire as an SMT solver. (experimental)
Elisabeth Lempa
parents: 80910
diff changeset
   203
  serialize = serialize sdtyp}
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   204
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   205
val _ = Theory.setup (Context.theory_map
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   206
  (setup_builtins #>
66551
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   207
   SMT_Translate.add_config (smtlibC, translate_config SMT_Util.First_Order) #>
4df6b0ae900d towards support for HO SMT-LIB
blanchet
parents: 66134
diff changeset
   208
   SMT_Translate.add_config (hosmtlibC, translate_config SMT_Util.Higher_Order)))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   209
57229
blanchet
parents: 57203
diff changeset
   210
end;