src/HOL/Tools/SMT/smtlib_interface.ML
author blanchet
Wed, 19 Nov 2014 10:31:15 +0100
changeset 59015 627a93f67182
parent 58482 7836013951e6
child 59021 b29281d6d1db
permissions -rw-r--r--
parse CVC4 unsat cores
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
56090
34bd10a9a2ad adapted to renamed ML files
blanchet
parents: 56088
diff changeset
    11
  val add_logic: int * (term list -> string option) -> Context.generic -> Context.generic
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    12
  val translate_config: Proof.context -> SMT_Translate.config
58482
7836013951e6 simplified and repaired veriT index handling code
blanchet
parents: 58429
diff changeset
    13
  val assert_name_of_index: int -> string
57203
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
    14
  val assert_index_of_name: string -> int
57704
c0da3fc313e3 Basic support for the SMT prover veriT.
fleury
parents: 57553
diff changeset
    15
  val assert_prefix : string
57229
blanchet
parents: 57203
diff changeset
    16
end;
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    17
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    18
structure SMTLIB_Interface: SMTLIB_INTERFACE =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    19
struct
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    20
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    21
val smtlibC = ["smtlib"]
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    22
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
(* builtins *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    25
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    26
local
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    27
  fun int_num _ i = SOME (string_of_int i)
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
  fun is_linear [t] = SMT_Util.is_number t
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    30
    | 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
    31
    | is_linear _ = false
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    32
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    33
  fun times _ _ ts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    34
    let val mk = Term.list_comb o pair @{const times (int)}
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    35
    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
    36
in
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    37
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    38
val setup_builtins =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    39
  fold (SMT_Builtin.add_builtin_typ smtlibC) [
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    40
    (@{typ bool}, K (SOME "Bool"), K (K NONE)),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    41
    (@{typ int}, K (SOME "Int"), int_num)] #>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    42
  fold (SMT_Builtin.add_builtin_fun' smtlibC) [
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    43
    (@{const True}, "true"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    44
    (@{const False}, "false"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    45
    (@{const Not}, "not"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    46
    (@{const HOL.conj}, "and"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    47
    (@{const HOL.disj}, "or"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    48
    (@{const HOL.implies}, "=>"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    49
    (@{const HOL.eq ('a)}, "="),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    50
    (@{const If ('a)}, "ite"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    51
    (@{const less (int)}, "<"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    52
    (@{const less_eq (int)}, "<="),
57711
caadd484dec6 Changing ~ into - for unuary minus (not supported by veriT)
fleury
parents: 57704
diff changeset
    53
    (@{const uminus (int)}, "-"),
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    54
    (@{const plus (int)}, "+"),
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    55
    (@{const minus (int)}, "-")] #>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    56
  SMT_Builtin.add_builtin_fun smtlibC
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    57
    (Term.dest_Const @{const times (int)}, times)
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
end
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
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    62
(* serialization *)
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    63
57238
blanchet
parents: 57229
diff changeset
    64
(** logic **)
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    65
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    66
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
    67
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    68
structure Logics = Generic_Data
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    69
(
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    70
  type T = (int * (term list -> string option)) list
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    71
  val empty = []
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    72
  val extend = I
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    73
  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
    74
)
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
fun add_logic pf = Logics.map (Ord_List.insert fst_int_ord pf)
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
fun choose_logic ctxt ts =
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    79
  let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    80
    fun choose [] = "AUFLIA"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    81
      | choose ((_, f) :: fs) = (case f 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
    82
  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
    83
    (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
    84
      "" => "" (* 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
    85
    | 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
    86
  end
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    87
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    88
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    89
(** serialization **)
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
fun var i = "?v" ^ string_of_int i
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    92
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    93
fun tree_of_sterm l (SMT_Translate.SVar i) = SMTLIB.Sym (var (l - i - 1))
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    94
  | tree_of_sterm _ (SMT_Translate.SApp (n, [])) = SMTLIB.Sym n
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    95
  | tree_of_sterm l (SMT_Translate.SApp (n, ts)) =
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    96
      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
    97
  | tree_of_sterm _ (SMT_Translate.SLet _) =
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
    98
      raise Fail "SMT-LIB: unsupported let expression"
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
    99
  | 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
   100
      let
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   101
        val l' = l + length ss
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   102
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   103
        fun quant_name SMT_Translate.SForall = "forall"
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   104
          | quant_name SMT_Translate.SExists = "exists"
56078
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
        fun gen_trees_of_pat keyword ps =
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   107
          [SMTLIB.Key keyword, (case map (tree_of_sterm l') ps of [t] => t | ts => SMTLIB.S ts)]
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   108
        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
   109
          | 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
   110
        fun tree_of_pats [] t = t
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   111
          | 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
   112
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   113
        val vs = map_index (fn (i, ty) =>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   114
          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
   115
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   116
        val body = t
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   117
          |> tree_of_sterm l'
57165
7b1bf424ec5f removed SMT weights -- their impact was very inconclusive anyway
blanchet
parents: 56091
diff changeset
   118
          |> tree_of_pats pats
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   119
      in
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   120
        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
   121
      end
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
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   124
fun sctrarg (sel, typ) = "(" ^ sel ^ " " ^ typ ^ ")"
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   125
fun sctr (name, args) = enclose "(" ")" (space_implode " " (name :: map sctrarg args))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   126
fun sdatatype (name, ctrs) = enclose "(" ")" (space_implode " " (name :: map sctr ctrs))
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   127
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   128
fun string_of_fun (f, (ss, s)) = f ^ " (" ^ space_implode " " ss ^ ") " ^ s
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   129
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   130
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
   131
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
   132
val assert_prefix = "a"
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
   133
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
   134
fun assert_name_of_index i = assert_prefix ^ string_of_int i
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
   135
fun assert_index_of_name s = the_default ~1 (Int.fromString (unprefix assert_prefix s))
171fa7d56c4f generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
blanchet
parents: 57165
diff changeset
   136
58429
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   137
fun sdtyp (fp, dtyps) =
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   138
  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
   139
    (space_implode "\n  " (map sdatatype dtyps)))
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   140
0b94858325a5 interleave (co)datatypes in the right order w.r.t. dependencies
blanchet
parents: 58361
diff changeset
   141
fun serialize smt_options comments {logic, sorts, dtyps, funcs} ts =
59015
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   142
  let
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   143
    val unsat_core = member (op =) smt_options (":produce-unsat-cores", "true")
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   144
  in
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   145
    Buffer.empty
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   146
    |> fold (Buffer.add o enclose "; " "\n") comments
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   147
    |> fold (fn (k, v) => Buffer.add ("(set-option " ^ k ^ " " ^ v ^ ")\n")) smt_options
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   148
    |> Buffer.add logic
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   149
    |> fold (Buffer.add o enclose "(declare-sort " " 0)\n") (sort fast_string_ord sorts)
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   150
    |> fold sdtyp (AList.coalesce (op =) dtyps)
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   151
    |> fold (Buffer.add o enclose "(declare-fun " ")\n" o string_of_fun)
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   152
        (sort (fast_string_ord o pairself fst) funcs)
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   153
    |> fold (fn (i, t) => Buffer.add (enclose "(assert " ")\n"
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   154
        (SMTLIB.str_of (named_sterm (assert_name_of_index i) (tree_of_sterm 0 t))))) (map_index I ts)
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   155
    |> Buffer.add "(check-sat)\n"
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   156
    |> Buffer.add (if unsat_core then "(get-unsat-core)\n" else "(get-proof)\n")
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   157
    |> Buffer.content
627a93f67182 parse CVC4 unsat cores
blanchet
parents: 58482
diff changeset
   158
  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
(* interface *)
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
fun translate_config ctxt = {
57238
blanchet
parents: 57229
diff changeset
   163
  logic = choose_logic ctxt,
58360
dee1fd1cc631 added interface for CVC4 extensions
blanchet
parents: 58061
diff changeset
   164
  fp_kinds = [],
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   165
  serialize = serialize}
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   166
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   167
val _ = Theory.setup (Context.theory_map
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   168
  (setup_builtins #>
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 57711
diff changeset
   169
   SMT_Translate.add_config (smtlibC, translate_config)))
56078
624faeda77b5 moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
blanchet
parents:
diff changeset
   170
57229
blanchet
parents: 57203
diff changeset
   171
end;