src/HOL/Tools/SMT/smtlib_interface.ML
author wenzelm
Sun, 30 Jan 2011 13:02:18 +0100
changeset 41648 6d736d983d5c
parent 41473 3717fc42ebe9
child 47155 ade3fc826af3
permissions -rw-r--r--
clarified example settings for Proof General;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Tools/SMT/smtlib_interface.ML
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     3
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     4
Interface to SMT solvers based on the SMT-LIB format.
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     5
*)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     6
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     7
signature SMTLIB_INTERFACE =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     8
sig
41124
1de17a2de5ad moved SMT classes and dictionary functions to SMT_Utils
boehmes
parents: 41072
diff changeset
     9
  val smtlibC: SMT_Utils.class
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    10
  val add_logic: int * (term list -> string option) -> Context.generic ->
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    11
    Context.generic
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
    12
  val translate_config: Proof.context -> SMT_Translate.config
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    13
  val setup: theory -> theory
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    14
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    15
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    16
structure SMTLIB_Interface: SMTLIB_INTERFACE =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    17
struct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    18
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    19
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    20
val smtlibC = ["smtlib"]
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    21
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    22
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    23
(* builtins *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    24
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    25
local
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    26
  fun int_num _ i = SOME (string_of_int i)
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    27
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
    28
  fun is_linear [t] = SMT_Utils.is_number t
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
    29
    | is_linear [t, u] = SMT_Utils.is_number t orelse SMT_Utils.is_number u
41280
a7de9d36f4f2 only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
boehmes
parents: 41127
diff changeset
    30
    | is_linear _ = false
a7de9d36f4f2 only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
boehmes
parents: 41127
diff changeset
    31
41281
679118e35378 removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents: 41280
diff changeset
    32
  fun times _ _ ts =
679118e35378 removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents: 41280
diff changeset
    33
    let val mk = Term.list_comb o pair @{const times (int)}
679118e35378 removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents: 41280
diff changeset
    34
    in if is_linear ts then SOME ("*", 2, ts, mk) else NONE end
41280
a7de9d36f4f2 only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
boehmes
parents: 41127
diff changeset
    35
41281
679118e35378 removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents: 41280
diff changeset
    36
  fun distinct _ T [t] =
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    37
        (case try HOLogic.dest_list t of
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
    38
          SOME (ts as _ :: _) =>
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
    39
            let
41281
679118e35378 removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents: 41280
diff changeset
    40
              val c = Const (@{const_name distinct}, T)
679118e35378 removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents: 41280
diff changeset
    41
              fun mk us = c $ HOLogic.mk_list T us
679118e35378 removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions);
boehmes
parents: 41280
diff changeset
    42
            in SOME ("distinct", length ts, ts, mk) end
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    43
        | _ => NONE)
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    44
    | distinct _ _ _ = NONE
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    45
in
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    46
41072
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
    47
val setup_builtins =
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
    48
  SMT_Builtin.add_builtin_typ smtlibC (@{typ int}, K (SOME "Int"), int_num) #>
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
    49
  fold (SMT_Builtin.add_builtin_fun' smtlibC) [
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    50
    (@{const True}, "true"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    51
    (@{const False}, "false"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    52
    (@{const Not}, "not"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    53
    (@{const HOL.conj}, "and"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    54
    (@{const HOL.disj}, "or"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    55
    (@{const HOL.implies}, "implies"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    56
    (@{const HOL.eq (bool)}, "iff"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    57
    (@{const HOL.eq ('a)}, "="),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    58
    (@{const If (bool)}, "if_then_else"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    59
    (@{const If ('a)}, "ite"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    60
    (@{const less (int)}, "<"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    61
    (@{const less_eq (int)}, "<="),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    62
    (@{const uminus (int)}, "~"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    63
    (@{const plus (int)}, "+"),
41280
a7de9d36f4f2 only linear occurrences of multiplication are treated as built-in (SMT solvers only support linear arithmetic in general);
boehmes
parents: 41127
diff changeset
    64
    (@{const minus (int)}, "-") ] #>
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
    65
  SMT_Builtin.add_builtin_fun smtlibC
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
    66
    (Term.dest_Const @{const times (int)}, times) #>
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
    67
  SMT_Builtin.add_builtin_fun smtlibC
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
    68
    (Term.dest_Const @{const distinct ('a)}, distinct)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    69
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    70
end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    71
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    72
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    73
(* serialization *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    74
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    75
(** header **)
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    76
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    77
fun fst_int_ord ((i1, _), (i2, _)) = int_ord (i1, i2)
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    78
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    79
structure Logics = Generic_Data
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    80
(
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    81
  type T = (int * (term list -> string option)) list
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    82
  val empty = []
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    83
  val extend = I
41473
3717fc42ebe9 Ord_List.merge convenience;
wenzelm
parents: 41426
diff changeset
    84
  fun merge data = Ord_List.merge fst_int_ord data
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    85
)
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    86
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    87
fun add_logic pf = Logics.map (Ord_List.insert fst_int_ord pf)
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    88
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    89
fun choose_logic ctxt ts =
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    90
  let
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    91
    fun choose [] = "AUFLIA"
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    92
      | choose ((_, f) :: fs) = (case f ts of SOME s => s | NONE => choose fs)
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    93
  in [":logic " ^ choose (Logics.get (Context.Proof ctxt))] end
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    94
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    95
41072
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
    96
(** serialization **)
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    97
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    98
val add = Buffer.add
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    99
fun sep f = add " " #> f
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   100
fun enclose l r f = sep (add l #> f #> add r)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   101
val par = enclose "(" ")"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   102
fun app n f = (fn [] => sep (add n) | xs => par (add n #> fold f xs))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   103
fun line f = f #> add "\n"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   104
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   105
fun var i = add "?v" #> add (string_of_int i)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   106
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   107
fun sterm l (SMT_Translate.SVar i) = sep (var (l - i - 1))
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   108
  | sterm l (SMT_Translate.SApp (n, ts)) = app n (sterm l) ts
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   109
  | sterm _ (SMT_Translate.SLet _) =
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   110
      raise Fail "SMT-LIB: unsupported let expression"
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   111
  | sterm l (SMT_Translate.SQua (q, ss, ps, w, t)) =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   112
      let
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   113
        fun quant SMT_Translate.SForall = add "forall"
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   114
          | quant SMT_Translate.SExists = add "exists"
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   115
        val vs = map_index (apfst (Integer.add l)) ss
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   116
        fun var_decl (i, s) = par (var i #> sep (add s))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   117
        val sub = sterm (l + length ss)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   118
        fun pat kind ts = sep (add kind #> enclose "{" " }" (fold sub ts))
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   119
        fun pats (SMT_Translate.SPat ts) = pat ":pat" ts
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   120
          | pats (SMT_Translate.SNoPat ts) = pat ":nopat" ts
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   121
        fun weight NONE = I
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   122
          | weight (SOME i) =
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   123
              sep (add ":weight { " #> add (string_of_int i) #> add " }")
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   124
      in
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   125
        par (quant q #> fold var_decl vs #> sub t #> fold pats ps #> weight w)
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   126
      end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   127
37155
e3f18cfc9829 sort signature in SMT-LIB output (improves sharing of SMT certificates: goals of the same logical structure are translated into equal SMT-LIB benchmarks)
boehmes
parents: 36936
diff changeset
   128
fun ssort sorts = sort fast_string_ord sorts
e3f18cfc9829 sort signature in SMT-LIB output (improves sharing of SMT certificates: goals of the same logical structure are translated into equal SMT-LIB benchmarks)
boehmes
parents: 36936
diff changeset
   129
fun fsort funcs = sort (prod_ord fast_string_ord (K EQUAL)) funcs
e3f18cfc9829 sort signature in SMT-LIB output (improves sharing of SMT certificates: goals of the same logical structure are translated into equal SMT-LIB benchmarks)
boehmes
parents: 36936
diff changeset
   130
39298
5aefb5bc8a93 added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents: 38864
diff changeset
   131
fun sdatatypes decls =
5aefb5bc8a93 added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents: 38864
diff changeset
   132
  let
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41328
diff changeset
   133
    fun con (n, []) = sep (add n)
39298
5aefb5bc8a93 added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents: 38864
diff changeset
   134
      | con (n, sels) = par (add n #>
41426
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41328
diff changeset
   135
          fold (fn (n, s) => par (add n #> sep (add s))) sels)
09615ed31f04 re-implemented support for datatypes (including records and typedefs);
boehmes
parents: 41328
diff changeset
   136
    fun dtyp (n, decl) = add n #> fold con decl
39298
5aefb5bc8a93 added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents: 38864
diff changeset
   137
  in line (add ":datatypes " #> par (fold (par o dtyp) decls)) end
5aefb5bc8a93 added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents: 38864
diff changeset
   138
5aefb5bc8a93 added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents: 38864
diff changeset
   139
fun serialize comments {header, sorts, dtyps, funcs} ts =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   140
  Buffer.empty
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   141
  |> line (add "(benchmark Isabelle")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   142
  |> line (add ":status unknown")
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   143
  |> fold (line o add) header
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   144
  |> length sorts > 0 ?
37155
e3f18cfc9829 sort signature in SMT-LIB output (improves sharing of SMT certificates: goals of the same logical structure are translated into equal SMT-LIB benchmarks)
boehmes
parents: 36936
diff changeset
   145
       line (add ":extrasorts" #> par (fold (sep o add) (ssort sorts)))
39298
5aefb5bc8a93 added preliminary support for SMT datatypes (for now restricted to tuples and lists); only the Z3 interface (in oracle mode) makes use of it, there is especially no Z3 proof reconstruction support for datatypes yet
boehmes
parents: 38864
diff changeset
   146
  |> fold sdatatypes dtyps
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   147
  |> length funcs > 0 ? (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   148
       line (add ":extrafuns" #> add " (") #>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   149
       fold (fn (f, (ss, s)) =>
37155
e3f18cfc9829 sort signature in SMT-LIB output (improves sharing of SMT certificates: goals of the same logical structure are translated into equal SMT-LIB benchmarks)
boehmes
parents: 36936
diff changeset
   150
         line (sep (app f (sep o add) (ss @ [s])))) (fsort funcs) #>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   151
       line (add ")"))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   152
  |> fold (fn t => line (add ":assumption" #> sterm 0 t)) ts
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   153
  |> line (add ":formula true)")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   154
  |> fold (fn str => line (add "; " #> add str)) comments
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   155
  |> Buffer.content
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   156
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   157
41072
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
   158
(* interface *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   159
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   160
fun translate_config ctxt = {
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   161
  prefixes = {
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   162
    sort_prefix = "S",
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   163
    func_prefix = "f"},
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   164
  header = choose_logic ctxt,
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   165
  is_fol = true,
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   166
  has_datatypes = false,
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   167
  serialize = serialize}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   168
41127
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   169
val setup = Context.theory_map (
2ea84c8535c6 re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
boehmes
parents: 41126
diff changeset
   170
  setup_builtins #>
41328
6792a5c92a58 avoid ML structure aliases (especially single-letter abbreviations)
boehmes
parents: 41281
diff changeset
   171
  SMT_Translate.add_config (smtlibC, translate_config))
41072
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
   172
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   173
end