src/HOL/Tools/SMT/smt_translate.ML
author boehmes
Mon, 06 Dec 2010 15:38:02 +0100
changeset 41057 8dbc951a291c
parent 40697 c3979dd80a50
child 41059 d2b1fc1b8e19
permissions -rw-r--r--
tuned
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/smt_translate.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
Translate theorems into an SMT intermediate format and serialize them.
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 SMT_TRANSLATE =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     8
sig
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
     9
  (* intermediate term structure *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    10
  datatype squant = SForall | SExists
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    11
  datatype 'a spattern = SPat of 'a list | SNoPat of 'a list
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    12
  datatype sterm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    13
    SVar of int |
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    14
    SApp of string * sterm list |
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    15
    SLet of string * sterm * sterm |
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
    16
    SQua of squant * string list * sterm spattern list * int option * sterm
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    17
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    18
  (* configuration options *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    19
  type prefixes = {sort_prefix: string, func_prefix: string}
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    20
  type header = Proof.context -> term list -> string list
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    21
  type strict = {
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    22
    is_builtin_conn: string * typ -> bool,
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    23
    is_builtin_pred: Proof.context -> string * typ -> bool,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    24
    is_builtin_distinct: bool}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    25
  type builtins = {
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    26
    builtin_typ: Proof.context -> typ -> string option,
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    27
    builtin_num: Proof.context -> typ -> int -> string option,
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    28
    builtin_fun: Proof.context -> string * typ -> term list ->
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: 37124
diff changeset
    29
      (string * term list) option,
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: 37124
diff changeset
    30
    has_datatypes: bool }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    31
  type sign = {
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    32
    header: string list,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    33
    sorts: string list,
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: 37124
diff changeset
    34
    dtyps: (string * (string * (string * string) list) list) list list,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    35
    funcs: (string * (string list * string)) list }
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    36
  type config = {
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    37
    prefixes: prefixes,
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    38
    header: header,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    39
    strict: strict option,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    40
    builtins: builtins,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    41
    serialize: string list -> sign -> sterm list -> string }
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    42
  type recon = {
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    43
    typs: typ Symtab.table,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    44
    terms: term Symtab.table,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    45
    unfolds: thm list,
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: 39535
diff changeset
    46
    assms: (int * thm) list }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    47
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: 39535
diff changeset
    48
  val translate: config -> Proof.context -> string list -> (int * thm) list ->
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    49
    string * recon
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    50
end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    51
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    52
structure SMT_Translate: SMT_TRANSLATE =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    53
struct
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    54
40663
e080c9e68752 share and use more utility functions;
boehmes
parents: 40579
diff changeset
    55
structure U = SMT_Utils
e080c9e68752 share and use more utility functions;
boehmes
parents: 40579
diff changeset
    56
e080c9e68752 share and use more utility functions;
boehmes
parents: 40579
diff changeset
    57
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    58
(* intermediate term structure *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    59
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    60
datatype squant = SForall | SExists
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    61
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    62
datatype 'a spattern = SPat of 'a list | SNoPat of 'a list
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    63
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    64
datatype sterm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    65
  SVar of int |
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    66
  SApp of string * sterm list |
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    67
  SLet of string * sterm * sterm |
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
    68
  SQua of squant * string list * sterm spattern list * int option * sterm
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    69
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    70
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    71
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    72
(* configuration options *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    73
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    74
type prefixes = {sort_prefix: string, func_prefix: string}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    75
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    76
type header = Proof.context -> term list -> string list
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    77
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    78
type strict = {
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    79
  is_builtin_conn: string * typ -> bool,
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    80
  is_builtin_pred: Proof.context -> string * typ -> bool,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    81
  is_builtin_distinct: bool}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    82
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    83
type builtins = {
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    84
  builtin_typ: Proof.context -> typ -> string option,
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    85
  builtin_num: Proof.context -> typ -> int -> string option,
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    86
  builtin_fun: Proof.context -> string * typ -> term list ->
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: 37124
diff changeset
    87
    (string * term list) option,
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: 37124
diff changeset
    88
  has_datatypes: bool }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    89
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    90
type sign = {
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    91
  header: string list,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    92
  sorts: string list,
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: 37124
diff changeset
    93
  dtyps: (string * (string * (string * string) list) list) list list,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    94
  funcs: (string * (string list * string)) list }
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    95
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    96
type config = {
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    97
  prefixes: prefixes,
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
    98
  header: header,
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
    99
  strict: strict option,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   100
  builtins: builtins,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   101
  serialize: string list -> sign -> sterm list -> string }
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   102
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   103
type recon = {
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   104
  typs: typ Symtab.table,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   105
  terms: term Symtab.table,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   106
  unfolds: thm list,
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: 39535
diff changeset
   107
  assms: (int * thm) list }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   108
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   109
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   110
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   111
(* utility functions *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   112
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   113
val quantifier = (fn
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   114
    @{const_name All} => SOME SForall
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   115
  | @{const_name Ex} => SOME SExists
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   116
  | _ => NONE)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   117
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   118
fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   119
      if q = qname then group_quant qname (T :: Ts) u else (Ts, t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   120
  | group_quant _ Ts t = (Ts, t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   121
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   122
fun dest_weight (@{const SMT.weight} $ w $ t) =
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   123
      (SOME (snd (HOLogic.dest_number w)), t)
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   124
  | dest_weight t = (NONE, t)
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   125
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   126
fun dest_pat (Const (@{const_name pat}, _) $ t) = (t, true)
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   127
  | dest_pat (Const (@{const_name nopat}, _) $ t) = (t, false)
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   128
  | dest_pat t = raise TERM ("dest_pat", [t])
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   129
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   130
fun dest_pats [] = I
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   131
  | dest_pats ts =
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   132
      (case map dest_pat ts |> split_list ||> distinct (op =) of
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   133
        (ps, [true]) => cons (SPat ps)
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   134
      | (ps, [false]) => cons (SNoPat ps)
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   135
      | _ => raise TERM ("dest_pats", ts))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   136
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40274
diff changeset
   137
fun dest_trigger (@{const trigger} $ tl $ t) =
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   138
      (rev (fold (dest_pats o HOLogic.dest_list) (HOLogic.dest_list tl) []), t)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   139
  | dest_trigger t = ([], t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   140
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   141
fun dest_quant qn T t = quantifier qn |> Option.map (fn q =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   142
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   143
    val (Ts, u) = group_quant qn [T] t
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   144
    val (ps, p) = dest_trigger u
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   145
    val (w, b) = dest_weight p
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   146
  in (q, rev Ts, ps, w, b) end)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   147
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   148
fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   149
  | fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   150
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   151
fun prop_of thm = HOLogic.dest_Trueprop (Thm.prop_of thm)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   152
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   153
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   154
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   155
(* enforce a strict separation between formulas and terms *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   156
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   157
val term_eq_rewr = @{lemma "term_eq x y == x = y" by (simp add: term_eq_def)}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   158
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   159
val term_bool = @{lemma "~(term_eq True False)" by (simp add: term_eq_def)}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   160
val term_bool' = Simplifier.rewrite_rule [term_eq_rewr] term_bool
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   161
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   162
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   163
val needs_rewrite = Thm.prop_of #> Term.exists_subterm (fn
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   164
    Const (@{const_name Let}, _) => true
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40274
diff changeset
   165
  | @{const HOL.eq (bool)} $ _ $ @{const True} => true
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40274
diff changeset
   166
  | Const (@{const_name If}, _) $ _ $ @{const True} $ @{const False} => true
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   167
  | _ => false)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   168
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   169
val rewrite_rules = [
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   170
  Let_def,
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   171
  @{lemma "P = True == P" by (rule eq_reflection) simp},
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   172
  @{lemma "if P then True else False == P" by (rule eq_reflection) simp}]
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   173
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   174
fun rewrite ctxt = Simplifier.full_rewrite
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   175
  (Simplifier.context ctxt empty_ss addsimps rewrite_rules)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   176
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   177
fun normalize ctxt thm =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   178
  if needs_rewrite thm then Conv.fconv_rule (rewrite ctxt) thm else thm
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   179
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   180
val unfold_rules = term_eq_rewr :: rewrite_rules
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   181
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   182
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   183
val revert_types =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   184
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   185
    fun revert @{typ prop} = @{typ bool}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   186
      | revert (Type (n, Ts)) = Type (n, map revert Ts)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   187
      | revert T = T
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   188
  in Term.map_types revert end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   189
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   190
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   191
fun strictify {is_builtin_conn, is_builtin_pred, is_builtin_distinct} ctxt =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   192
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   193
    fun is_builtin_conn' (@{const_name True}, _) = false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   194
      | is_builtin_conn' (@{const_name False}, _) = false
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   195
      | is_builtin_conn' c = is_builtin_conn c
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   196
41057
boehmes
parents: 40697
diff changeset
   197
    fun is_builtin_pred' _ (@{const_name distinct}, _) [t] =
40681
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   198
          is_builtin_distinct andalso can HOLogic.dest_list t
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   199
      | is_builtin_pred' ctxt c _ = is_builtin_pred ctxt c
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   200
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   201
    val propT = @{typ prop} and boolT = @{typ bool}
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   202
    val as_propT = (fn @{typ bool} => propT | T => T)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   203
    fun mapTs f g = Term.strip_type #> (fn (Ts, T) => map f Ts ---> g T)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   204
    fun conn (n, T) = (n, mapTs as_propT as_propT T)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   205
    fun pred (n, T) = (n, mapTs I as_propT T)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   206
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40274
diff changeset
   207
    val term_eq = @{const HOL.eq (bool)} |> Term.dest_Const |> pred
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40274
diff changeset
   208
    fun as_term t = Const term_eq $ t $ @{const True}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   209
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   210
    val if_term = Const (@{const_name If}, [propT, boolT, boolT] ---> boolT)
40579
98ebd2300823 use the const antiquotation for constants (this checks that the constant is declared, whereas the more general term antiquotation treats undeclared names as free variable)
boehmes
parents: 40274
diff changeset
   211
    fun wrap_in_if t = if_term $ t $ @{const True} $ @{const False}
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   212
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   213
    fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   214
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   215
    fun in_term t =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   216
      (case Term.strip_comb t of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   217
        (c as Const (@{const_name If}, _), [t1, t2, t3]) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   218
          c $ in_form t1 $ in_term t2 $ in_term t3
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   219
      | (h as Const c, ts) =>
40681
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   220
          if is_builtin_conn' (conn c) orelse is_builtin_pred' ctxt (pred c) ts
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   221
          then wrap_in_if (in_form t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   222
          else Term.list_comb (h, map in_term ts)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   223
      | (h as Free _, ts) => Term.list_comb (h, map in_term ts)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   224
      | _ => t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   225
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   226
    and in_weight ((c as @{const SMT.weight}) $ w $ t) = c $ w $ in_form t
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   227
      | in_weight t = in_form t 
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   228
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   229
    and in_pat ((c as Const (@{const_name pat}, _)) $ t) = c $ in_term t
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   230
      | in_pat ((c as Const (@{const_name nopat}, _)) $ t) = c $ in_term t
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   231
      | in_pat t = raise TERM ("in_pat", [t])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   232
37124
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   233
    and in_pats ps =
fe22fc54b876 hide constants and types introduced by SMT,
boehmes
parents: 36899
diff changeset
   234
      in_list @{typ "pattern list"} (in_list @{typ pattern} in_pat) ps
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   235
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   236
    and in_trig ((c as @{const trigger}) $ p $ t) = c $ in_pats p $ in_weight t
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   237
      | in_trig t = in_weight t
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   238
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   239
    and in_form t =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   240
      (case Term.strip_comb t of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   241
        (q as Const (qn, _), [Abs (n, T, t')]) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   242
          if is_some (quantifier qn) then q $ Abs (n, T, in_trig t')
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   243
          else as_term (in_term t)
40681
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   244
      | (Const (c as (@{const_name distinct}, T)), [t']) =>
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   245
          if is_builtin_distinct andalso can HOLogic.dest_list t' then
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   246
            Const (pred c) $ in_list T in_term t'
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   247
          else as_term (in_term t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   248
      | (Const c, ts) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   249
          if is_builtin_conn (conn c)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   250
          then Term.list_comb (Const (conn c), map in_form ts)
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   251
          else if is_builtin_pred ctxt (pred c)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   252
          then Term.list_comb (Const (pred c), map in_term ts)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   253
          else as_term (in_term t)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   254
      | _ => as_term (in_term t))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   255
  in
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: 39535
diff changeset
   256
    map (apsnd (normalize ctxt)) #> (fn irules =>
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: 39535
diff changeset
   257
    ((unfold_rules, (~1, term_bool') :: irules),
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: 39535
diff changeset
   258
     map (in_form o prop_of o snd) ((~1, term_bool) :: irules)))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   259
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   260
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   261
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   262
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   263
(* translation from Isabelle terms into SMT intermediate terms *)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   264
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: 37124
diff changeset
   265
val empty_context = (1, Typtab.empty, [], 1, Termtab.empty)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   266
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: 37124
diff changeset
   267
fun make_sign header (_, typs, dtyps, _, terms) = {
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   268
  header = header,
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: 37124
diff changeset
   269
  sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [],
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: 37124
diff changeset
   270
  funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms [],
39435
5d18f4c00c07 reverse order of datatype declarations so that declarations only depend on already declared datatypes
boehmes
parents: 39298
diff changeset
   271
  dtyps = rev dtyps }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   272
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: 37124
diff changeset
   273
fun make_recon (unfolds, assms) (_, typs, _, _, terms) = {
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: 37124
diff changeset
   274
  typs = Symtab.make (map (apfst fst o swap) (Typtab.dest typs)),
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: 37124
diff changeset
   275
    (*FIXME: don't drop the datatype information! *)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   276
  terms = Symtab.make (map (fn (t, (n, _)) => (n, t)) (Termtab.dest terms)),
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   277
  unfolds = unfolds,
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   278
  assms = assms }
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   279
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   280
fun string_of_index pre i = pre ^ string_of_int i
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   281
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: 37124
diff changeset
   282
fun new_typ sort_prefix proper T (Tidx, typs, dtyps, idx, terms) =
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: 37124
diff changeset
   283
  let val s = string_of_index sort_prefix Tidx
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: 37124
diff changeset
   284
  in (s, (Tidx+1, Typtab.update (T, (s, proper)) typs, dtyps, idx, terms)) 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: 37124
diff changeset
   285
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: 37124
diff changeset
   286
fun lookup_typ (_, typs, _, _, _) = Typtab.lookup typs
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   287
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: 37124
diff changeset
   288
fun fresh_typ T f cx =
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: 37124
diff changeset
   289
  (case lookup_typ cx T of
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: 37124
diff changeset
   290
    SOME (s, _) => (s, cx)
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: 37124
diff changeset
   291
  | NONE => f T cx)
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: 37124
diff changeset
   292
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: 37124
diff changeset
   293
fun new_fun func_prefix t ss (Tidx, typs, dtyps, idx, terms) =
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: 37124
diff changeset
   294
  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: 37124
diff changeset
   295
    val f = string_of_index func_prefix idx
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: 37124
diff changeset
   296
    val terms' = Termtab.update (revert_types t, (f, ss)) terms
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: 37124
diff changeset
   297
  in (f, (Tidx, typs, dtyps, idx+1, terms')) 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: 37124
diff changeset
   298
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: 37124
diff changeset
   299
fun fresh_fun func_prefix t ss (cx as (_, _, _, _, terms)) =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   300
  (case Termtab.lookup terms t of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   301
    SOME (f, _) => (f, cx)
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: 37124
diff changeset
   302
  | NONE => new_fun func_prefix t ss cx)
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: 37124
diff changeset
   303
39483
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   304
fun mk_type (_, Tfs) (d as Datatype.DtTFree _) = the (AList.lookup (op =) Tfs d)
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   305
  | mk_type Ts (Datatype.DtType (n, ds)) = Type (n, map (mk_type Ts) ds)
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   306
  | mk_type (Tds, _) (Datatype.DtRec i) = nth Tds i
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: 37124
diff changeset
   307
39483
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   308
fun mk_selector ctxt Ts T n (i, d) =
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   309
  (case Datatype_Selectors.lookup_selector ctxt (n, i+1) of
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   310
    NONE => raise Fail ("missing selector for datatype constructor " ^ quote n)
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   311
  | SOME m => mk_type Ts d |> (fn U => (Const (m, T --> U), U)))
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   312
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   313
fun mk_constructor ctxt Ts T (n, args) =
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   314
  let val (sels, Us) = split_list (map_index (mk_selector ctxt Ts T n) args)
9f0e5684f04b add full support for datatypes to the SMT interface (only used by Z3 in oracle mode so far); added store to keep track of datatype selector functions
boehmes
parents: 39435
diff changeset
   315
  in (Const (n, Us ---> T), sels) end
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: 37124
diff changeset
   316
39535
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   317
fun lookup_datatype ctxt n Ts =
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   318
  if member (op =) [@{type_name bool}, @{type_name nat}] n then NONE
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   319
  else
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   320
    Datatype.get_info (ProofContext.theory_of ctxt) n
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   321
    |> Option.map (fn {descr, ...} =>
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   322
         let
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   323
           val Tds = map (fn (_, (tn, _, _)) => Type (tn, Ts))
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   324
             (sort (int_ord o pairself fst) descr)
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   325
           val Tfs = (case hd descr of (_, (_, tfs, _)) => tfs ~~ Ts)
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   326
         in
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   327
           descr |> map (fn (i, (_, _, cs)) =>
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   328
             (nth Tds i, map (mk_constructor ctxt (Tds, Tfs) (nth Tds i)) cs))
cd1bb7125d8a do not treat natural numbers as a datatype (natural numbers are considered an abstract type with a coercion to integers)
boehmes
parents: 39483
diff changeset
   329
         end)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   330
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: 39535
diff changeset
   331
fun relaxed irules = (([], irules), map (prop_of o snd) irules)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   332
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   333
fun with_context header f (ths, ts) =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   334
  let val (us, context) = fold_map f ts empty_context
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   335
  in ((make_sign (header ts) context, us), make_recon ths context) end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   336
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   337
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   338
fun translate {prefixes, strict, header, builtins, serialize} ctxt comments =
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   339
  let
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   340
    val {sort_prefix, func_prefix} = prefixes
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: 37124
diff changeset
   341
    val {builtin_typ, builtin_num, builtin_fun, has_datatypes} = builtins
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: 37124
diff changeset
   342
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: 37124
diff changeset
   343
    fun transT (T as TFree _) = fresh_typ T (new_typ sort_prefix true)
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: 37124
diff changeset
   344
      | transT (T as TVar _) = (fn _ => raise TYPE ("smt_translate", [T], []))
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: 37124
diff changeset
   345
      | transT (T as Type (n, Ts)) =
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: 37124
diff changeset
   346
          (case builtin_typ ctxt T of
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: 37124
diff changeset
   347
            SOME n => pair 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: 37124
diff changeset
   348
          | NONE => fresh_typ T (fn _ => fn cx =>
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: 37124
diff changeset
   349
              if not has_datatypes then new_typ sort_prefix true T cx
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: 37124
diff changeset
   350
              else
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: 37124
diff changeset
   351
                (case lookup_datatype ctxt n Ts of
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: 37124
diff changeset
   352
                  NONE => new_typ sort_prefix true T cx
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: 37124
diff changeset
   353
                | SOME dts =>
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: 37124
diff changeset
   354
                    let val cx' = new_dtyps dts cx 
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: 37124
diff changeset
   355
                    in (fst (the (lookup_typ cx' T)), cx') end)))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   356
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: 37124
diff changeset
   357
    and new_dtyps dts cx =
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: 37124
diff changeset
   358
      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: 37124
diff changeset
   359
        fun new_decl i t =
40663
e080c9e68752 share and use more utility functions;
boehmes
parents: 40579
diff changeset
   360
          let val (Ts, T) = U.dest_funT i (Term.fastype_of t)
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: 37124
diff changeset
   361
          in
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: 37124
diff changeset
   362
            fold_map transT Ts ##>> transT T ##>>
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: 37124
diff changeset
   363
            new_fun func_prefix t NONE #>> swap
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: 37124
diff changeset
   364
          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: 37124
diff changeset
   365
        fun new_dtyp_decl (con, 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: 37124
diff changeset
   366
          new_decl (length sels) con ##>> fold_map (new_decl 1) 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: 37124
diff changeset
   367
          (fn ((con', _), sels') => (con', map (apsnd snd) 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: 37124
diff changeset
   368
      in
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: 37124
diff changeset
   369
        cx
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: 37124
diff changeset
   370
        |> fold_map (new_typ sort_prefix false o fst) dts
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: 37124
diff changeset
   371
        ||>> fold_map (fold_map new_dtyp_decl o snd) dts
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: 37124
diff changeset
   372
        |-> (fn (ss, decls) => fn (Tidx, typs, dtyps, idx, terms) =>
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: 37124
diff changeset
   373
              (Tidx, typs, (ss ~~ decls) :: dtyps, idx, terms))
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: 37124
diff changeset
   374
      end
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   375
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   376
    fun app n ts = SApp (n, ts)
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   377
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   378
    fun trans t =
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   379
      (case Term.strip_comb t of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   380
        (Const (qn, _), [Abs (_, T, t1)]) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   381
          (case dest_quant qn T t1 of
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   382
            SOME (q, Ts, ps, w, b) =>
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   383
              fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>>
40664
e023788a91a1 added support for quantifier weight annotations
boehmes
parents: 40663
diff changeset
   384
              trans b #>> (fn ((Ts', ps'), b') => SQua (q, Ts', ps', w, b'))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   385
          | NONE => raise TERM ("intermediate", [t]))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   386
      | (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   387
          transT T ##>> trans t1 ##>> trans t2 #>>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   388
          (fn ((U, u1), u2) => SLet (U, u1, u2))
40681
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   389
      | (h as Const (c as (@{const_name distinct}, T)), ts) =>
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   390
          (case builtin_fun ctxt c ts of
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   391
            SOME (n, ts) => fold_map trans ts #>> app n
40681
872b08416fb4 be more precise: only treat constant 'distinct' applied to an explicit list as built-in
boehmes
parents: 40664
diff changeset
   392
          | NONE => transs h T ts)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   393
      | (h as Const (c as (_, T)), ts) =>
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   394
          (case try HOLogic.dest_number t of
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   395
            SOME (T, i) =>
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   396
              (case builtin_num ctxt T i of
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   397
                SOME n => pair (SApp (n, []))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   398
              | NONE => transs t T [])
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   399
          | NONE =>
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   400
              (case builtin_fun ctxt c ts of
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   401
                SOME (n, ts') => fold_map trans ts' #>> app n
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   402
              | NONE => transs h T ts))
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   403
      | (h as Free (_, T), ts) => transs h T ts
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   404
      | (Bound i, []) => pair (SVar i)
40697
c3979dd80a50 eta-reduce on the fly to prevent an exception
blanchet
parents: 40681
diff changeset
   405
      | (Abs (_, _, t1 $ Bound 0), []) =>
c3979dd80a50 eta-reduce on the fly to prevent an exception
blanchet
parents: 40681
diff changeset
   406
        if not (loose_bvar1 (t1, 0)) then trans t1 (* eta-reduce on the fly *)
c3979dd80a50 eta-reduce on the fly to prevent an exception
blanchet
parents: 40681
diff changeset
   407
        else raise TERM ("smt_translate", [t])
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: 37124
diff changeset
   408
      | _ => raise TERM ("smt_translate", [t]))
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   409
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   410
    and transs t T ts =
40663
e080c9e68752 share and use more utility functions;
boehmes
parents: 40579
diff changeset
   411
      let val (Us, U) = U.dest_funT (length ts) T
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   412
      in
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   413
        fold_map transT Us ##>> transT U #-> (fn Up =>
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: 37124
diff changeset
   414
        fresh_fun func_prefix t (SOME Up) ##>> fold_map trans ts #>> SApp)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   415
      end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   416
  in
36899
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   417
    (case strict of SOME strct => strictify strct ctxt | NONE => relaxed) #>
bcd6fce5bf06 layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents: 36898
diff changeset
   418
    with_context (header ctxt) trans #>> uncurry (serialize comments)
36898
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   419
  end
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   420
8e55aa1306c5 integrated SMT into the HOL image
boehmes
parents:
diff changeset
   421
end