src/HOL/Tools/SMT/smtlib_interface.ML
author boehmes
Wed, 08 Dec 2010 08:33:02 +0100
changeset 41072 9f9bc1bdacef
parent 41059 d2b1fc1b8e19
child 41124 1de17a2de5ad
permissions -rw-r--r--
be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
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
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
     9
  val smtlibC: SMT_Config.class
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
41072
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
    12
  val interface: SMT_Solver.interface
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
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    19
structure B = SMT_Builtin
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    20
structure N = SMT_Normalize
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    21
structure T = SMT_Translate
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
val smtlibC = ["smtlib"]
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    24
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    25
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    26
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    27
(* facts about uninterpreted constants *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    28
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    29
infix 2 ??
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39687
diff changeset
    30
fun (ex ?? f) irules = irules |> exists (ex o Thm.prop_of o snd) irules ? f
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    31
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    32
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    33
(** pairs **)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    34
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    35
val pair_rules = [@{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    36
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37155
diff changeset
    37
val pair_type = (fn Type (@{type_name Product_Type.prod}, _) => true | _ => false)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    38
val exists_pair_type = Term.exists_type (Term.exists_subtype pair_type)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    39
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39687
diff changeset
    40
val add_pair_rules = exists_pair_type ?? append (map (pair ~1) pair_rules)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    41
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    42
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    43
(** function update **)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    44
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    45
val fun_upd_rules = [@{thm fun_upd_same}, @{thm fun_upd_apply}]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    46
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    47
val is_fun_upd = (fn Const (@{const_name fun_upd}, _) => true | _ => false)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    48
val exists_fun_upd = Term.exists_subterm is_fun_upd
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    49
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39687
diff changeset
    50
val add_fun_upd_rules = exists_fun_upd ?? append (map (pair ~1) fun_upd_rules)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    51
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    52
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    53
(** abs/min/max **)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    54
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    55
val exists_abs_min_max = Term.exists_subterm (fn
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    56
    Const (@{const_name abs}, _) => true
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    57
  | Const (@{const_name min}, _) => true
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    58
  | Const (@{const_name max}, _) => true
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    59
  | _ => false)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    60
37786
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37678
diff changeset
    61
val unfold_abs_conv = Conv.rewr_conv (mk_meta_eq @{thm abs_if})
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37678
diff changeset
    62
val unfold_min_conv = Conv.rewr_conv (mk_meta_eq @{thm min_def})
4eb98849c5c0 fixed handling of Ball/Bex: turn equalities into meta-equalities for the rewriting conversions;
boehmes
parents: 37678
diff changeset
    63
val unfold_max_conv = Conv.rewr_conv (mk_meta_eq @{thm max_def})
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    64
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    65
fun expand_conv cv = N.eta_expand_conv (K cv)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    66
fun expand2_conv cv = N.eta_expand_conv (N.eta_expand_conv (K cv))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    67
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    68
fun unfold_def_conv ctxt ct =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    69
  (case Thm.term_of ct of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    70
    Const (@{const_name abs}, _) $ _ => unfold_abs_conv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    71
  | Const (@{const_name abs}, _) => expand_conv unfold_abs_conv ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    72
  | Const (@{const_name min}, _) $ _ $ _ => unfold_min_conv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    73
  | Const (@{const_name min}, _) $ _ => expand_conv unfold_min_conv ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    74
  | Const (@{const_name min}, _) => expand2_conv unfold_min_conv ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    75
  | Const (@{const_name max}, _) $ _ $ _ => unfold_max_conv
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    76
  | Const (@{const_name max}, _) $ _ => expand_conv unfold_max_conv ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    77
  | Const (@{const_name max}, _) => expand2_conv unfold_max_conv ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    78
  | _ => Conv.all_conv) ct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    79
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    80
fun unfold_abs_min_max_defs ctxt thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    81
  if exists_abs_min_max (Thm.prop_of thm)
36936
c52d1c130898 incorporated further conversions and conversionals, after some minor tuning;
wenzelm
parents: 36899
diff changeset
    82
  then Conv.fconv_rule (Conv.top_conv unfold_def_conv ctxt) thm
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    83
  else thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    84
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    85
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    86
(** include additional facts **)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    87
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39687
diff changeset
    88
fun extra_norm has_datatypes irules ctxt =
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39687
diff changeset
    89
  irules
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
    90
  |> not has_datatypes ? add_pair_rules
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    91
  |> add_fun_upd_rules
40161
539d07b00e5f keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents: 39687
diff changeset
    92
  |> map (apsnd (unfold_abs_min_max_defs ctxt))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    93
  |> rpair ctxt
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    94
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    95
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    96
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    97
(* builtins *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    98
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
    99
local
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   100
  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
   101
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   102
  fun distinct _ _ [t] =
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   103
        (case try HOLogic.dest_list t of
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   104
          SOME (ts as _ :: _) => SOME ("distinct", ts)
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   105
        | _ => NONE)
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   106
    | distinct _ _ _ = NONE
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   107
in
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   108
41072
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
   109
val setup_builtins =
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   110
  B.add_builtin_typ smtlibC (@{typ int}, K (SOME "Int"), int_num) #>
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   111
  fold (B.add_builtin_fun' smtlibC) [
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   112
    (@{const True}, "true"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   113
    (@{const False}, "false"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   114
 (* FIXME: we do not test if these constants are fully applied! *)
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   115
    (@{const Not}, "not"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   116
    (@{const HOL.conj}, "and"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   117
    (@{const HOL.disj}, "or"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   118
    (@{const HOL.implies}, "implies"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   119
    (@{const HOL.eq (bool)}, "iff"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   120
    (@{const HOL.eq ('a)}, "="),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   121
    (@{const If (bool)}, "if_then_else"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   122
    (@{const If ('a)}, "ite"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   123
    (@{const less (int)}, "<"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   124
    (@{const less_eq (int)}, "<="),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   125
    (@{const uminus (int)}, "~"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   126
    (@{const plus (int)}, "+"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   127
    (@{const minus (int)}, "-"),
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   128
    (@{const times (int)}, "*") ] #>
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   129
  B.add_builtin_fun smtlibC (Term.dest_Const @{const distinct ('a)}, distinct)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   130
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   131
end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   132
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   133
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   134
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   135
(* serialization *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   136
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   137
(** header **)
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   138
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   139
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
   140
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   141
structure Logics = Generic_Data
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   142
(
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   143
  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
   144
  val empty = []
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   145
  val extend = I
39687
4e9b6ada3a21 modernized structure Ord_List;
wenzelm
parents: 39298
diff changeset
   146
  fun merge (bs1, bs2) = Ord_List.union fst_int_ord bs2 bs1
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   147
)
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   148
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   149
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
   150
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   151
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
   152
  let
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   153
    fun choose [] = "AUFLIA"
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   154
      | 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
   155
  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
   156
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
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
(** serialization **)
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   159
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   160
val add = Buffer.add
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   161
fun sep f = add " " #> f
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   162
fun enclose l r f = sep (add l #> f #> add r)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   163
val par = enclose "(" ")"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   164
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
   165
fun line f = f #> add "\n"
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   166
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   167
fun var i = add "?v" #> add (string_of_int i)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   168
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   169
fun sterm l (T.SVar i) = sep (var (l - i - 1))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   170
  | sterm l (T.SApp (n, ts)) = app n (sterm l) ts
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   171
  | sterm _ (T.SLet _) = raise Fail "SMT-LIB: unsupported let expression"
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   172
  | sterm l (T.SQua (q, ss, ps, w, t)) =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   173
      let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   174
        val quant = add o (fn T.SForall => "forall" | T.SExists => "exists")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   175
        val vs = map_index (apfst (Integer.add l)) ss
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   176
        fun var_decl (i, s) = par (var i #> sep (add s))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   177
        val sub = sterm (l + length ss)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   178
        fun pat kind ts = sep (add kind #> enclose "{" " }" (fold sub ts))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   179
        fun pats (T.SPat ts) = pat ":pat" ts
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   180
          | pats (T.SNoPat ts) = pat ":nopat" ts
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   181
        fun weight NONE = I
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   182
          | weight (SOME i) =
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   183
              sep (add ":weight { " #> add (string_of_int i) #> add " }")
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   184
      in
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40274
diff changeset
   185
        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
   186
      end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   187
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
   188
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
   189
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
   190
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
   191
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
   192
  let
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
   193
    fun con (n, []) = add n
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
   194
      | con (n, sels) = par (add n #>
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
   195
          fold (fn (n, s) => sep (par (add n #> sep (add s)))) sels)
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
   196
    fun dtyp (n, decl) = add n #> fold (sep o con) decl
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
   197
  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
   198
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
   199
fun serialize comments {header, sorts, dtyps, funcs} ts =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   200
  Buffer.empty
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   201
  |> line (add "(benchmark Isabelle")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   202
  |> 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
   203
  |> fold (line o add) header
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   204
  |> 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
   205
       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
   206
  |> fold sdatatypes dtyps
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   207
  |> length funcs > 0 ? (
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   208
       line (add ":extrafuns" #> add " (") #>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   209
       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
   210
         line (sep (app f (sep o add) (ss @ [s])))) (fsort funcs) #>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   211
       line (add ")"))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   212
  |> fold (fn t => line (add ":assumption" #> sterm 0 t)) ts
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   213
  |> line (add ":formula true)")
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   214
  |> fold (fn str => line (add "; " #> add str)) comments
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   215
  |> Buffer.content
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   216
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   217
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   218
41072
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
   219
(* interface *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   220
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   221
val interface = {
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   222
  class = smtlibC,
40162
7f58a9a843c2 joined setup of SMT solvers in one place; turned Z3-specific options into SMT options (renamed configuration options from z3_* to smt_*); more detailed SMT exception; improved SMT filter interface
boehmes
parents: 40161
diff changeset
   223
  extra_norm = extra_norm,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   224
  translate = {
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   225
    prefixes = {
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   226
      sort_prefix = "S",
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   227
      func_prefix = "f"},
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   228
    header = choose_logic,
41059
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   229
    is_fol = true,
d2b1fc1b8e19 centralized handling of built-in types and constants;
boehmes
parents: 41057
diff changeset
   230
    has_datatypes = false,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   231
    serialize = serialize}}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   232
41072
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
   233
val setup = Context.theory_map setup_builtins
9f9bc1bdacef be more flexible: store SMT built-in symbols in generic contexts (not in theory contexts)
boehmes
parents: 41059
diff changeset
   234
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   235
end