re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure);
authorboehmes
Wed, 15 Dec 2010 10:12:44 +0100
changeset 41127 2ea84c8535c6
parent 41126 e0bd443c0fdd
child 41128 bb2fa5c13d1a
re-implemented eta-expansion, lambda-lifting, and explicit application on terms (exploiting the control over the term structure); abolished SMT interface concept in favor of solver classes (now also the translation configuration is stored in the context); proof reconstruction is now expected to return a theorem stating False (and hence needs to discharge all hypothetical definitions); built-in functions carry additionally their arity and their most general type; slightly generalized the definition of fun_app
src/HOL/SMT.thy
src/HOL/Tools/SMT/smt_builtin.ML
src/HOL/Tools/SMT/smt_setup_solvers.ML
src/HOL/Tools/SMT/smt_solver.ML
src/HOL/Tools/SMT/smt_translate.ML
src/HOL/Tools/SMT/smt_utils.ML
src/HOL/Tools/SMT/smtlib_interface.ML
src/HOL/Tools/SMT/z3_interface.ML
src/HOL/Tools/SMT/z3_proof_reconstruction.ML
src/HOL/Word/Tools/smt_word.ML
--- a/src/HOL/SMT.thy	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/SMT.thy	Wed Dec 15 10:12:44 2010 +0100
@@ -91,7 +91,7 @@
 following constant.
 *}
 
-definition fun_app where "fun_app f x = f x"
+definition fun_app where "fun_app f = f"
 
 text {*
 Some solvers support a theory of arrays which can be used to encode
--- a/src/HOL/Tools/SMT/smt_builtin.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_builtin.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -17,15 +17,15 @@
   val is_builtin_typ_ext: Proof.context -> typ -> bool
 
   (*built-in numbers*)
-  val builtin_num: Proof.context -> term -> string option
+  val builtin_num: Proof.context -> term -> (string * typ) option
   val is_builtin_num: Proof.context -> term -> bool
   val is_builtin_num_ext: Proof.context -> term -> bool
 
   (*built-in functions*)
   type 'a bfun = Proof.context -> typ -> term list -> 'a
   val add_builtin_fun: SMT_Utils.class ->
-    (string * typ) * (string * term list) option bfun -> Context.generic ->
-    Context.generic
+    (string * typ) * (((string * int) * typ) * term list * typ) option bfun ->
+    Context.generic -> Context.generic
   val add_builtin_fun': SMT_Utils.class -> term * string -> Context.generic ->
     Context.generic
   val add_builtin_fun_ext: (string * typ) * bool bfun -> Context.generic ->
@@ -33,10 +33,8 @@
   val add_builtin_fun_ext': string * typ -> Context.generic -> Context.generic
   val add_builtin_fun_ext'': string -> Context.generic -> Context.generic
   val builtin_fun: Proof.context -> string * typ -> term list ->
-    (string * term list) option
+    (((string * int) * typ) * term list * typ) option
   val is_builtin_fun: Proof.context -> string * typ -> term list -> bool
-  val is_builtin_pred: Proof.context -> string * typ -> term list -> bool
-  val is_builtin_conn: Proof.context -> string * typ -> term list -> bool
   val is_builtin_fun_ext: Proof.context -> string * typ -> term list -> bool
   val is_builtin_ext: Proof.context -> string * typ -> term list -> bool
 end
@@ -131,7 +129,7 @@
     NONE => NONE
   | SOME (T, i) =>
       (case lookup_builtin_typ ctxt T of
-        SOME (_, Int (_, g)) => g T i
+        SOME (_, Int (_, g)) => g T i |> Option.map (rpair T)
       | _ => NONE))
 
 val is_builtin_num = is_some oo builtin_num
@@ -148,7 +146,8 @@
 
 structure Builtin_Funcs = Generic_Data
 (
-  type T = (bool bfun, (string * term list) option bfun) btab
+  type T =
+    (bool bfun, (((string * int) * typ) * term list * typ) option bfun) btab
   val empty = Symtab.empty
   val extend = I
   val merge = merge_btab
@@ -158,7 +157,10 @@
   Builtin_Funcs.map (insert_btab cs n T (Int f))
 
 fun add_builtin_fun' cs (t, n) =
-  add_builtin_fun cs (Term.dest_Const t, fn _ => fn _ => SOME o pair n)
+  let
+    val T = Term.fastype_of t
+    fun bfun _ U ts = SOME (((n, length (Term.binder_types T)), U), ts, T)
+  in add_builtin_fun cs (Term.dest_Const t, bfun) end
 
 fun add_builtin_fun_ext ((n, T), f) =
   Builtin_Funcs.map (insert_btab U.basicC n T (Ext f))
@@ -180,18 +182,6 @@
 
 fun is_builtin_fun ctxt c ts = is_some (builtin_fun ctxt c ts)
 
-fun is_special_builtin_fun pred ctxt (c as (_, T)) ts =
-  (case lookup_builtin_fun ctxt c of
-    SOME (U, Int f) => pred U andalso is_some (f ctxt T ts)
-  | _ => false)
-
-fun is_pred_type T = Term.body_type T = @{typ bool}
-fun is_conn_type T =
-  forall (equal @{typ bool}) (Term.body_type T :: Term.binder_types T)
-
-fun is_builtin_pred ctxt = is_special_builtin_fun is_pred_type ctxt
-fun is_builtin_conn ctxt = is_special_builtin_fun is_conn_type ctxt
-
 fun is_builtin_fun_ext ctxt (c as (_, T)) ts =
   (case lookup_builtin_fun ctxt c of
     SOME (_, Int f) => is_some (f ctxt T ts)
--- a/src/HOL/Tools/SMT/smt_setup_solvers.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_setup_solvers.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -38,7 +38,7 @@
   options = (fn ctxt => [
     "-seed", string_of_int (Config.get ctxt SMT_Config.random_seed),
     "-lang", "smtlib", "-output-lang", "presentation"]),
-  interface = SMTLIB_Interface.interface,
+  class = SMTLIB_Interface.smtlibC,
   outcome =
     on_first_line (outcome_of "Unsatisfiable." "Satisfiable." "Unknown."),
   cex_parser = NONE,
@@ -55,7 +55,7 @@
   options = (fn ctxt => [
     "--rand-seed", string_of_int (Config.get ctxt SMT_Config.random_seed),
     "--smtlib"]),
-  interface = SMTLIB_Interface.interface,
+  class = SMTLIB_Interface.smtlibC,
   outcome = on_first_line (outcome_of "unsat" "sat" "unknown"),
   cex_parser = NONE,
   reconstruct = NONE,
@@ -89,7 +89,7 @@
   env_var = "Z3_SOLVER",
   is_remote = true,
   options = z3_options,
-  interface = Z3_Interface.interface,
+  class = Z3_Interface.smtlib_z3C,
   outcome = z3_on_last_line,
   cex_parser = SOME Z3_Model.parse_counterex,
   reconstruct = SOME Z3_Proof_Reconstruction.reconstruct,
--- a/src/HOL/Tools/SMT/smt_solver.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_solver.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -7,21 +7,18 @@
 signature SMT_SOLVER =
 sig
   (*configuration*)
-  type interface = {
-    class: SMT_Utils.class,
-    translate: SMT_Translate.config }
   datatype outcome = Unsat | Sat | Unknown
   type solver_config = {
     name: string,
     env_var: string,
     is_remote: bool,
     options: Proof.context -> string list,
-    interface: interface,
+    class: SMT_Utils.class,
     outcome: string -> string list -> outcome * string list,
     cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
       term list * term list) option,
     reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
-      (int list * thm) * Proof.context) option,
+      int list * thm) option,
     default_max_relevant: int }
 
   (*registry*)
@@ -57,10 +54,6 @@
 
 (* configuration *)
 
-type interface = {
-  class: SMT_Utils.class,
-  translate: SMT_Translate.config }
-
 datatype outcome = Unsat | Sat | Unknown
 
 type solver_config = {
@@ -68,12 +61,12 @@
   env_var: string,
   is_remote: bool,
   options: Proof.context -> string list,
-  interface: interface,
+  class: SMT_Utils.class,
   outcome: string -> string list -> outcome * string list,
   cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
     term list * term list) option,
   reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
-    (int list * thm) * Proof.context) option,
+    int list * thm) option,
   default_max_relevant: int }
 
 
@@ -163,7 +156,7 @@
 fun trace_assms ctxt = C.trace_msg ctxt (Pretty.string_of o
   Pretty.big_list "Assertions:" o map (Display.pretty_thm ctxt o snd))
 
-fun trace_recon_data ctxt ({typs, terms, ...} : SMT_Translate.recon) =
+fun trace_recon_data ({context=ctxt, typs, terms, ...} : SMT_Translate.recon) =
   let
     fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
     fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
@@ -175,32 +168,20 @@
         Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()
   end
 
-fun invoke translate_config name cmd options ithms ctxt =
+fun invoke name cmd options ithms ctxt =
   let
     val args = C.solver_options_of ctxt @ options ctxt
     val comments = ("solver: " ^ name) ::
       ("timeout: " ^ string_of_real (Config.get ctxt C.timeout)) ::
       ("random seed: " ^ string_of_int (Config.get ctxt C.random_seed)) ::
       "arguments:" :: args
-  in
-    ithms
-    |> tap (trace_assms ctxt)
-    |> SMT_Translate.translate translate_config ctxt comments
-    ||> tap (trace_recon_data ctxt)
-    |>> run_solver ctxt cmd args
-    |> rpair ctxt
-  end
 
-fun discharge_definitions thm =
-  if Thm.nprems_of thm = 0 then thm
-  else discharge_definitions (@{thm reflexive} RS thm)
-
-fun set_has_datatypes with_datatypes translate =
-  let val {prefixes, header, is_fol, has_datatypes, serialize} = translate
-  in
-   {prefixes=prefixes, header=header, is_fol=is_fol,
-    has_datatypes=has_datatypes andalso with_datatypes, serialize=serialize}
-  end
+    val (str, recon as {context=ctxt', ...}) =
+      ithms
+      |> tap (trace_assms ctxt)
+      |> SMT_Translate.translate ctxt comments
+      ||> tap trace_recon_data
+  in (run_solver ctxt' cmd args str, recon) end
 
 fun trace_assumptions ctxt iwthms idxs =
   let
@@ -227,26 +208,21 @@
   env_var: string,
   is_remote: bool,
   options: Proof.context -> string list,
-  interface: interface,
-  reconstruct: string list * SMT_Translate.recon -> Proof.context ->
-    (int list * thm) * Proof.context,
+  reconstruct: Proof.context -> string list * SMT_Translate.recon ->
+    int list * thm,
   default_max_relevant: int }
 
 fun gen_solver name (info : solver_info) rm ctxt iwthms =
   let
-    val {env_var, is_remote, options, interface, reconstruct, ...} = info
-    val {translate, ...} = interface
-    val translate' = set_has_datatypes (Config.get ctxt C.datatypes) translate
+    val {env_var, is_remote, options, reconstruct, ...} = info
     val cmd = (rm, env_var, is_remote, name)
   in
     SMT_Normalize.normalize ctxt iwthms
     |> rpair ctxt
     |-> SMT_Monomorph.monomorph
-    |-> invoke translate' name cmd options
-    |-> reconstruct
-    |-> (fn (idxs, thm) => fn ctxt' => thm
-    |> singleton (ProofContext.export ctxt' ctxt)
-    |> discharge_definitions
+    |-> invoke name cmd options
+    |> reconstruct ctxt
+    |> (fn (idxs, thm) => thm
     |> tap (fn _ => trace_assumptions ctxt iwthms idxs)
     |> pair idxs)
   end
@@ -261,12 +237,13 @@
 )
 
 local
-  fun finish outcome cex_parser reconstruct ocl (output, recon) ctxt =
+  fun finish outcome cex_parser reconstruct ocl outer_ctxt
+      (output, (recon as {context=ctxt, ...} : SMT_Translate.recon)) =
     (case outcome output of
       (Unsat, ls) =>
         if not (Config.get ctxt C.oracle) andalso is_some reconstruct
-        then the reconstruct ctxt recon ls
-        else (([], ocl ()), ctxt)
+        then the reconstruct outer_ctxt recon ls
+        else ([], ocl ())
     | (result, ls) =>
         let
           val (ts, us) =
@@ -283,14 +260,12 @@
 
 fun add_solver cfg =
   let
-    val {name, env_var, is_remote, options, interface, outcome, cex_parser,
+    val {name, env_var, is_remote, options, class, outcome, cex_parser,
       reconstruct, default_max_relevant} = cfg
-    val {class, ...} = interface
 
     fun core_oracle () = cfalse
 
     fun solver ocl = { env_var=env_var, is_remote=is_remote, options=options,
-      interface=interface,
       reconstruct=finish (outcome name) cex_parser reconstruct ocl,
       default_max_relevant=default_max_relevant }
   in
--- a/src/HOL/Tools/SMT/smt_translate.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_translate.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -15,7 +15,7 @@
     SLet of string * sterm * sterm |
     SQua of squant * string list * sterm spattern list * int option * sterm
 
-  (*configuration options*)
+  (*translation configuration*)
   type prefixes = {sort_prefix: string, func_prefix: string}
   type sign = {
     header: string list,
@@ -24,17 +24,21 @@
     funcs: (string * (string list * string)) list }
   type config = {
     prefixes: prefixes,
-    header: Proof.context -> term list -> string list,
+    header: term list -> string list,
     is_fol: bool,
     has_datatypes: bool,
     serialize: string list -> sign -> sterm list -> string }
   type recon = {
+    context: Proof.context,
     typs: typ Symtab.table,
     terms: term Symtab.table,
-    unfolds: thm list,
+    rewrite_rules: thm list,
     assms: (int * thm) list }
 
-  val translate: config -> Proof.context -> string list -> (int * thm) list ->
+  (*translation*)
+  val add_config: SMT_Utils.class * (Proof.context -> config) ->
+    Context.generic -> Context.generic 
+  val translate: Proof.context -> string list -> (int * thm) list ->
     string * recon
 end
 
@@ -59,7 +63,7 @@
 
 
 
-(* configuration options *)
+(* translation configuration *)
 
 type prefixes = {sort_prefix: string, func_prefix: string}
 
@@ -71,20 +75,416 @@
 
 type config = {
   prefixes: prefixes,
-  header: Proof.context -> term list -> string list,
+  header: term list -> string list,
   is_fol: bool,
   has_datatypes: bool,
   serialize: string list -> sign -> sterm list -> string }
 
 type recon = {
+  context: Proof.context,
   typs: typ Symtab.table,
   terms: term Symtab.table,
-  unfolds: thm list,
+  rewrite_rules: thm list,
   assms: (int * thm) list }
 
 
 
-(* utility functions *)
+(* translation context *)
+
+fun make_tr_context {sort_prefix, func_prefix} =
+  (sort_prefix, 1, Typtab.empty, func_prefix, 1, Termtab.empty)
+
+fun string_of_index pre i = pre ^ string_of_int i
+
+fun add_typ T proper (cx as (sp, Tidx, typs, fp, idx, terms)) =
+  (case Typtab.lookup typs T of
+    SOME (n, _) => (n, cx)
+  | NONE =>
+      let
+        val n = string_of_index sp Tidx
+        val typs' = Typtab.update (T, (n, proper)) typs
+      in (n, (sp, Tidx+1, typs', fp, idx, terms)) end)
+
+fun add_fun t sort (cx as (sp, Tidx, typs, fp, idx, terms)) =
+  (case Termtab.lookup terms t of
+    SOME (n, _) => (n, cx)
+  | NONE => 
+      let
+        val n = string_of_index fp idx
+        val terms' = Termtab.update (t, (n, sort)) terms
+      in (n, (sp, Tidx, typs, fp, idx+1, terms')) end)
+
+fun sign_of header dtyps (_, _, typs, _, _, terms) = {
+  header = header,
+  sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [],
+  dtyps = dtyps,
+  funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms []}
+
+fun recon_of ctxt rules thms ithms revertT revert (_, _, typs, _, _, terms) =
+  let
+    fun add_typ (T, (n, _)) = Symtab.update (n, revertT T)
+    val typs' = Typtab.fold add_typ typs Symtab.empty
+
+    fun add_fun (t, (n, _)) = Symtab.update (n, revert t)
+    val terms' = Termtab.fold add_fun terms Symtab.empty
+
+    val assms = map (pair ~1) thms @ ithms
+  in
+    {context=ctxt, typs=typs', terms=terms', rewrite_rules=rules, assms=assms}
+  end
+
+
+
+(* preprocessing *)
+
+(** mark built-in constants as Var **)
+
+fun mark_builtins ctxt =
+  let
+    (*
+      Note: schematic terms cannot occur anymore in terms at this stage.
+    *)
+    fun mark t =
+      (case Term.strip_comb t of
+        (u as Const (@{const_name If}, _), ts) => marks u ts
+      | (u as Const c, ts) =>
+          (case B.builtin_num ctxt t of
+            SOME (n, T) =>
+              let val v = ((n, 0), T)
+              in Vartab.update v #> pair (Var v) end
+          | NONE =>
+              (case B.builtin_fun ctxt c ts of
+                SOME ((ni, T), us, U) =>
+                  Vartab.update (ni, U) #> marks (Var (ni, T)) us
+              | NONE => marks u ts))
+      | (Abs (n, T, u), ts) => mark u #-> (fn u' => marks (Abs (n, T, u')) ts)
+      | (u, ts) => marks u ts)
+ 
+    and marks t ts = fold_map mark ts #>> Term.list_comb o pair t
+
+  in (fn ts => swap (fold_map mark ts Vartab.empty)) end
+
+fun mark_builtins' ctxt t = hd (snd (mark_builtins ctxt [t]))
+
+
+(** FIXME **)
+
+local
+  (*
+    mark constructors and selectors as Vars (forcing eta-expansion),
+    add missing datatype selectors via hypothetical definitions,
+    also return necessary datatype and record theorems
+  *)
+in
+
+fun collect_datatypes_and_records (tr_context, ctxt) ts =
+  (([], tr_context, ctxt), ts)
+
+end
+
+
+(** eta-expand quantifiers, let expressions and built-ins *)
+
+local
+  fun eta T t = Abs (Name.uu, T, Term.incr_boundvars 1 t $ Bound 0)
+
+  fun exp T = eta (Term.domain_type (Term.domain_type T))
+
+  fun exp2 T q =
+    let val U = Term.domain_type T
+    in Abs (Name.uu, U, q $ eta (Term.domain_type U) (Bound 0)) end
+
+  fun exp2' T l =
+    let val (U1, U2) = Term.dest_funT T ||> Term.domain_type
+    in Abs (Name.uu, U1, eta U2 (l $ Bound 0)) end
+
+  fun expf t i T ts =
+    let val Ts = U.dest_funT i T |> fst |> drop (length ts)
+    in Term.list_comb (t, ts) |> fold_rev eta Ts end
+
+  fun expand ((q as Const (@{const_name All}, _)) $ Abs a) = q $ abs_expand a
+    | expand ((q as Const (@{const_name All}, T)) $ t) = q $ exp T t
+    | expand (q as Const (@{const_name All}, T)) = exp2 T q
+    | expand ((q as Const (@{const_name Ex}, _)) $ Abs a) = q $ abs_expand a
+    | expand ((q as Const (@{const_name Ex}, T)) $ t) = q $ exp T t
+    | expand (q as Const (@{const_name Ex}, T)) = exp2 T q
+    | expand ((l as Const (@{const_name Let}, _)) $ t $ Abs a) =
+        l $ expand t $ abs_expand a
+    | expand ((l as Const (@{const_name Let}, T)) $ t $ u) =
+        l $ expand t $ exp (Term.range_type T) u
+    | expand ((l as Const (@{const_name Let}, T)) $ t) = exp2 T (l $ expand t)
+    | expand (l as Const (@{const_name Let}, T)) = exp2' T l
+    | expand (Abs a) = abs_expand a
+    | expand t =
+        (case Term.strip_comb t of
+          (u as Const (@{const_name If}, T), ts) => expf u 3 T (map expand ts)
+        | (u as Var ((_, i), T), ts) => expf u i T (map expand ts)
+        | (u, ts) => Term.list_comb (u, map expand ts))
+
+  and abs_expand (n, T, t) = Abs (n, T, expand t)
+in
+
+val eta_expand = map expand
+
+end
+
+
+(** lambda-lifting **)
+
+local
+  fun mk_def Ts T lhs rhs =
+    let
+      val eq = HOLogic.eq_const T $ lhs $ rhs
+      val trigger =
+        [[Const (@{const_name SMT.pat}, T --> @{typ SMT.pattern}) $ lhs]]
+        |> map (HOLogic.mk_list @{typ SMT.pattern})
+        |> HOLogic.mk_list @{typ "SMT.pattern list"}
+      fun mk_all T t = HOLogic.all_const T $ Abs (Name.uu, T, t)
+    in fold mk_all Ts (@{const SMT.trigger} $ trigger $ eq) end
+
+  fun replace_lambda Us Ts t (cx as (defs, ctxt)) =
+    let
+      val T = Term.fastype_of1 (Us @ Ts, t)
+      val lev = length Us
+      val bs = sort int_ord (Term.add_loose_bnos (t, lev, []))
+      val bss = map_index (fn (i, j) => (j, Integer.add lev i)) bs
+      val norm = perhaps (AList.lookup (op =) bss)
+      val t' = Term.map_aterms (fn Bound i => Bound (norm i) | t => t) t
+      val Ts' = map (nth Ts) bs
+
+      fun mk_abs U u = Abs (Name.uu, U, u)
+      val abs_rhs = fold mk_abs Ts' (fold mk_abs Us t')
+    in
+      (case Termtab.lookup defs abs_rhs of
+        SOME (f, _) => (Term.list_comb (f, map Bound bs), cx)
+      | NONE =>
+          let
+            val (n, ctxt') =
+              yield_singleton Variable.variant_fixes Name.uu ctxt
+            val f = Free (n, rev Ts' ---> (rev Us ---> T))
+            fun mk_bapp i t = t $ Bound i
+            val lhs =
+              f
+              |> fold_rev (mk_bapp o snd) bss
+              |> fold_rev mk_bapp (0 upto (length Us - 1))
+            val def = mk_def (Us @ Ts') T lhs t'
+          in (f, (Termtab.update (abs_rhs, (f, def)) defs, ctxt')) end)
+    end
+
+  fun dest_abs Ts (Abs (_, T, t)) = dest_abs (T :: Ts) t
+    | dest_abs Ts t = (Ts, t)
+
+  fun traverse Ts t =
+    (case t of
+      (q as Const (@{const_name All}, _)) $ Abs a =>
+        abs_traverse Ts a #>> (fn a' => q $ Abs a')
+    | (q as Const (@{const_name Ex}, _)) $ Abs a =>
+        abs_traverse Ts a #>> (fn a' => q $ Abs a')
+    | (l as Const (@{const_name Let}, _)) $ u $ Abs a =>
+        traverse Ts u ##>> abs_traverse Ts a #>>
+        (fn (u', a') => l $ u' $ Abs a')
+    | Abs _ =>
+        let val (Us, u) = dest_abs [] t
+        in traverse (Us @ Ts) u #-> replace_lambda Us Ts end
+    | u1 $ u2 => traverse Ts u1 ##>> traverse Ts u2 #>> (op $)
+    | _ => pair t)
+
+  and abs_traverse Ts (n, T, t) = traverse (T::Ts) t #>> (fn t' => (n, T, t'))
+in
+
+fun lift_lambdas ctxt ts =
+  (Termtab.empty, ctxt)
+  |> fold_map (traverse []) ts
+  |> (fn (us, (defs, ctxt')) =>
+       (ctxt', Termtab.fold (cons o snd o snd) defs us))
+
+end
+
+
+(** introduce explicit applications **)
+
+local
+  (*
+    Make application explicit for functions with varying number of arguments.
+  *)
+
+  fun add t ts =
+    Termtab.map_default (t, []) (Ord_List.insert int_ord (length ts))
+
+  fun collect t =
+    (case Term.strip_comb t of
+      (u as Const _, ts) => add u ts #> fold collect ts
+    | (u as Free _, ts) => add u ts #> fold collect ts
+    | (Abs (_, _, u), ts) => collect u #> fold collect ts
+    | (_, ts) => fold collect ts)
+
+  fun app ts (t, T) =
+    let val f = Const (@{const_name SMT.fun_app}, T --> T)
+    in (Term.list_comb (f $ t, ts), snd (U.dest_funT (length ts) T)) end 
+
+  fun appl _ _ [] = fst
+    | appl _ [] ts = fst o app ts
+    | appl i (k :: ks) ts =
+        let val (ts1, ts2) = chop (k - i) ts
+        in appl k ks ts2 o app ts1 end
+
+  fun appl0 [_] ts (t, _) = Term.list_comb (t, ts)
+    | appl0 (0 :: ks) ts tT = appl 0 ks ts tT
+    | appl0 ks ts tT = appl 0 ks ts tT
+
+  fun apply terms T t ts = appl0 (Termtab.lookup_list terms t) ts (t, T)
+
+  fun get_arities i t =
+    (case Term.strip_comb t of
+      (Bound j, ts) =>
+        (if i = j then Ord_List.insert int_ord (length ts) else I) #>
+        fold (get_arities i) ts
+    | (Abs (_, _, u), ts) => get_arities (i+1) u #> fold (get_arities i) ts
+    | (_, ts) => fold (get_arities i) ts)
+in
+
+fun intro_explicit_application ts =
+  let
+    val terms = fold collect ts Termtab.empty
+
+    fun traverse (env as (arities, Ts)) t =
+      (case Term.strip_comb t of
+        (u as Const (_, T), ts) => apply terms T u (map (traverse env) ts)
+      | (u as Free (_, T), ts) => apply terms T u (map (traverse env) ts)
+      | (u as Bound i, ts) =>
+          appl0 (nth arities i) (map (traverse env) ts) (u, nth Ts i)
+      | (Abs (n, T, u), ts) =>
+          let val env' = (get_arities 0 u [] :: arities, T :: Ts)
+          in traverses env (Abs (n, T, traverse env' u)) ts end
+      | (u, ts) => traverses env u ts)
+    and traverses env t ts = Term.list_comb (t, map (traverse env) ts)
+  in map (traverse ([], [])) ts end
+
+val fun_app_eq = mk_meta_eq @{thm SMT.fun_app_def}
+
+end
+
+
+(** map HOL formulas to FOL formulas (i.e., separate formulas froms terms) **)
+
+val tboolT = @{typ SMT.term_bool}
+val term_true = Const (@{const_name True}, tboolT)
+val term_false = Const (@{const_name False}, tboolT)
+
+val term_bool = @{lemma "True ~= False" by simp}
+val term_bool_prop =
+  let
+    fun replace @{const HOL.eq (bool)} = @{const HOL.eq (SMT.term_bool)}
+      | replace @{const True} = term_true
+      | replace @{const False} = term_false
+      | replace t = t
+  in
+    Term.map_aterms replace (HOLogic.dest_Trueprop (Thm.prop_of term_bool))
+  end
+
+val fol_rules = [
+  Let_def,
+  @{lemma "P = True == P" by (rule eq_reflection) simp},
+  @{lemma "if P then True else False == P" by (rule eq_reflection) simp}]
+
+fun reduce_let (Const (@{const_name Let}, _) $ t $ u) =
+      reduce_let (Term.betapply (u, t))
+  | reduce_let (t $ u) = reduce_let t $ reduce_let u
+  | reduce_let (Abs (n, T, t)) = Abs (n, T, reduce_let t)
+  | reduce_let t = t
+
+fun is_pred_type NONE = false
+  | is_pred_type (SOME T) = (Term.body_type T = @{typ bool})
+
+fun is_conn_type NONE = false
+  | is_conn_type (SOME T) =
+      forall (equal @{typ bool}) (Term.body_type T :: Term.binder_types T)
+
+fun revert_typ @{typ SMT.term_bool} = @{typ bool}
+  | revert_typ (Type (n, Ts)) = Type (n, map revert_typ Ts)
+  | revert_typ T = T
+
+val revert_types = Term.map_types revert_typ
+
+fun folify ctxt builtins =
+  let
+    fun as_term t = @{const HOL.eq (SMT.term_bool)} $ t $ term_true
+
+    fun as_tbool @{typ bool} = tboolT
+      | as_tbool (Type (n, Ts)) = Type (n, map as_tbool Ts)
+      | as_tbool T = T
+    fun mapTs f g i = U.dest_funT i #> (fn (Ts, T) => map f Ts ---> g T)
+    fun predT i = mapTs as_tbool I i
+    fun funcT i = mapTs as_tbool as_tbool i
+    fun func i (n, T) = (n, funcT i T)
+
+    fun map_ifT T = T |> Term.dest_funT ||> funcT 2 |> (op -->)
+    val if_term = @{const If (bool)} |> Term.dest_Const ||> map_ifT |> Const
+    fun wrap_in_if t = if_term $ t $ term_true $ term_false
+
+    fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t))
+
+    fun in_term t =
+      (case Term.strip_comb t of
+        (Const (n as @{const_name If}, T), [t1, t2, t3]) =>
+          Const (n, map_ifT T) $ in_form t1 $ in_term t2 $ in_term t3
+      | (Const (@{const_name HOL.eq}, _), _) => wrap_in_if (in_form t)
+      | (Var (ni as (_, i), T), ts) =>
+          let val U = Vartab.lookup builtins ni
+          in
+            if is_conn_type U orelse is_pred_type U then wrap_in_if (in_form t)
+            else Term.list_comb (Var (ni, funcT i T), map in_term ts)
+          end
+      | (Const c, ts) =>
+          Term.list_comb (Const (func (length ts) c), map in_term ts)
+      | (Free c, ts) =>
+          Term.list_comb (Free (func (length ts) c), map in_term ts)
+      | _ => t)
+
+    and in_weight ((c as @{const SMT.weight}) $ w $ t) = c $ w $ in_form t
+      | in_weight t = in_form t 
+
+    and in_pat (Const (c as (@{const_name pat}, _)) $ t) =
+          Const (func 1 c) $ in_term t
+      | in_pat (Const (c as (@{const_name nopat}, _)) $ t) =
+          Const (func 1 c) $ in_term t
+      | in_pat t = raise TERM ("bad pattern", [t])
+
+    and in_pats ps =
+      in_list @{typ "pattern list"} (in_list @{typ pattern} in_pat) ps
+
+    and in_trig ((c as @{const trigger}) $ p $ t) = c $ in_pats p $ in_weight t
+      | in_trig t = in_weight t
+
+    and in_form t =
+      (case Term.strip_comb t of
+        (q as Const (qn, _), [Abs (n, T, u)]) =>
+          if member (op =) [@{const_name All}, @{const_name Ex}] qn then
+            q $ Abs (n, as_tbool T, in_trig u)
+          else as_term (in_term t)
+      | (u as Const (@{const_name If}, _), ts) =>
+          Term.list_comb (u, map in_form ts)
+      | (b as @{const HOL.eq (bool)}, ts) => Term.list_comb (b, map in_form ts)
+      | (Const (n as @{const_name HOL.eq}, T), ts) =>
+          Term.list_comb (Const (n, predT 2 T), map in_term ts)
+      | (b as Var (ni as (_, i), T), ts) =>
+          if is_conn_type (Vartab.lookup builtins ni) then
+            Term.list_comb (b, map in_form ts)
+          else if is_pred_type (Vartab.lookup builtins ni) then
+            Term.list_comb (Var (ni, predT i T), map in_term ts)
+          else as_term (in_term t)
+      | _ => as_term (in_term t))
+  in
+    map (reduce_let #> in_form) #>
+    cons (mark_builtins' ctxt term_bool_prop) #>
+    pair (fol_rules, [term_bool])
+  end
+
+
+
+(* translation into intermediate format *)
+
+(** utility functions **)
 
 val quantifier = (fn
     @{const_name All} => SOME SForall
@@ -101,14 +501,14 @@
 
 fun dest_pat (Const (@{const_name pat}, _) $ t) = (t, true)
   | dest_pat (Const (@{const_name nopat}, _) $ t) = (t, false)
-  | dest_pat t = raise TERM ("dest_pat", [t])
+  | dest_pat t = raise TERM ("bad pattern", [t])
 
 fun dest_pats [] = I
   | dest_pats ts =
       (case map dest_pat ts |> split_list ||> distinct (op =) of
         (ps, [true]) => cons (SPat ps)
       | (ps, [false]) => cons (SNoPat ps)
-      | _ => raise TERM ("dest_pats", ts))
+      | _ => raise TERM ("bad multi-pattern", ts))
 
 fun dest_trigger (@{const trigger} $ tl $ t) =
       (rev (fold (dest_pats o HOLogic.dest_list) (HOLogic.dest_list tl) []), t)
@@ -124,233 +524,19 @@
 fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat
   | fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat
 
-fun prop_of thm = HOLogic.dest_Trueprop (Thm.prop_of thm)
 
-
-
-(* map HOL formulas to FOL formulas (i.e., separate formulas froms terms) *)
-
-val tboolT = @{typ SMT.term_bool}
-val term_true = Const (@{const_name True}, tboolT)
-val term_false = Const (@{const_name False}, tboolT)
-
-val term_bool = @{lemma "True ~= False" by simp}
-val term_bool_prop =
-  let
-    fun replace @{const HOL.eq (bool)} = @{const HOL.eq (SMT.term_bool)}
-      | replace @{const True} = term_true
-      | replace @{const False} = term_false
-      | replace t = t
-  in Term.map_aterms replace (prop_of term_bool) end
-
-val needs_rewrite = Thm.prop_of #> Term.exists_subterm (fn
-    Const (@{const_name Let}, _) => true
-  | @{const HOL.eq (bool)} $ _ $ @{const True} => true
-  | Const (@{const_name If}, _) $ _ $ @{const True} $ @{const False} => true
-  | _ => false)
-
-val rewrite_rules = [
-  Let_def,
-  @{lemma "P = True == P" by (rule eq_reflection) simp},
-  @{lemma "if P then True else False == P" by (rule eq_reflection) simp}]
-
-fun rewrite ctxt ct =
-  Conv.top_sweep_conv (fn ctxt' =>
-    Conv.rewrs_conv rewrite_rules then_conv rewrite ctxt') ctxt ct
-
-fun normalize ctxt thm =
-  if needs_rewrite thm then Conv.fconv_rule (rewrite ctxt) thm else thm
-
-fun revert_typ @{typ SMT.term_bool} = @{typ bool}
-  | revert_typ (Type (n, Ts)) = Type (n, map revert_typ Ts)
-  | revert_typ T = T
-
-val revert_types = Term.map_types revert_typ
-
-fun folify ctxt =
-  let
-    fun is_builtin_conn (@{const_name True}, _) _ = false
-      | is_builtin_conn (@{const_name False}, _) _ = false
-      | is_builtin_conn c ts = B.is_builtin_conn ctxt c ts
-
-    fun as_term t = @{const HOL.eq (SMT.term_bool)} $ t $ term_true
-
-    fun as_tbool @{typ bool} = tboolT
-      | as_tbool (Type (n, Ts)) = Type (n, map as_tbool Ts)
-      | as_tbool T = T
-    fun mapTs f g = Term.strip_type #> (fn (Ts, T) => map f Ts ---> g T)
-    fun predT T = mapTs as_tbool I T
-    fun funcT T = mapTs as_tbool as_tbool T
-    fun func (n, T) = Const (n, funcT T)
-
-    fun map_ifT T = T |> Term.dest_funT ||> funcT |> (op -->)
-    val if_term = @{const If (bool)} |> Term.dest_Const ||> map_ifT |> Const
-    fun wrap_in_if t = if_term $ t $ term_true $ term_false
-
-    fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t))
-
-    fun in_term t =
-      (case Term.strip_comb t of
-        (Const (c as @{const_name If}, T), [t1, t2, t3]) =>
-          Const (c, map_ifT T) $ in_form t1 $ in_term t2 $ in_term t3
-      | (Const c, ts) =>
-          if is_builtin_conn c ts orelse B.is_builtin_pred ctxt c ts
-          then wrap_in_if (in_form t)
-          else Term.list_comb (func c, map in_term ts)
-      | (Free (n, T), ts) => Term.list_comb (Free (n, funcT T), map in_term ts)
-      | _ => t)
-
-    and in_weight ((c as @{const SMT.weight}) $ w $ t) = c $ w $ in_form t
-      | in_weight t = in_form t 
-
-    and in_pat (Const (c as (@{const_name pat}, _)) $ t) = func c $ in_term t
-      | in_pat (Const (c as (@{const_name nopat}, _)) $ t) = func c $ in_term t
-      | in_pat t = raise TERM ("in_pat", [t])
-
-    and in_pats ps =
-      in_list @{typ "pattern list"} (in_list @{typ pattern} in_pat) ps
-
-    and in_trig ((c as @{const trigger}) $ p $ t) = c $ in_pats p $ in_weight t
-      | in_trig t = in_weight t
+(** translation from Isabelle terms into SMT intermediate terms **)
 
-    and in_form t =
-      (case Term.strip_comb t of
-        (q as Const (qn, _), [Abs (n, T, t')]) =>
-          if is_some (quantifier qn) then q $ Abs (n, as_tbool T, in_trig t')
-          else as_term (in_term t)
-      | (Const (c as (n as @{const_name distinct}, T)), [t']) =>
-          if B.is_builtin_fun ctxt c [t'] then
-            Const (n, predT T) $ in_list T in_term t'
-          else as_term (in_term t)
-      | (Const (c as (n, T)), ts) =>
-          if B.is_builtin_conn ctxt c ts
-          then Term.list_comb (Const c, map in_form ts)
-          else if B.is_builtin_pred ctxt c ts
-          then Term.list_comb (Const (n, predT T), map in_term ts)
-          else as_term (in_term t)
-      | _ => as_term (in_term t))
-  in
-    map (apsnd (normalize ctxt)) #> (fn irules =>
-    ((rewrite_rules, (~1, term_bool) :: irules),
-     term_bool_prop :: map (in_form o prop_of o snd) irules))
-  end
-
-
-
-(* translation from Isabelle terms into SMT intermediate terms *)
-
-val empty_context = (1, Typtab.empty, [], 1, Termtab.empty)
-
-fun make_sign header (_, typs, dtyps, _, terms) = {
-  header = header,
-  sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [],
-  funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms [],
-  dtyps = rev dtyps }
-
-fun make_recon (unfolds, assms) (_, typs, _, _, terms) = {
-  typs = Symtab.make (map (apfst fst o swap) (Typtab.dest typs)),
-    (*FIXME: don't drop the datatype information! *)
-  terms = Symtab.make (map (fn (t, (n, _)) => (n, t)) (Termtab.dest terms)),
-  unfolds = unfolds,
-  assms = assms }
-
-fun string_of_index pre i = pre ^ string_of_int i
-
-fun new_typ sort_prefix proper T (Tidx, typs, dtyps, idx, terms) =
+fun intermediate header dtyps ctxt ts trx =
   let
-    val s = string_of_index sort_prefix Tidx
-    val U = revert_typ T
-  in (s, (Tidx+1, Typtab.update (U, (s, proper)) typs, dtyps, idx, terms)) end
-
-fun lookup_typ (_, typs, _, _, _) = Typtab.lookup typs o revert_typ
-
-fun fresh_typ T f cx =
-  (case lookup_typ cx T of
-    SOME (s, _) => (s, cx)
-  | NONE => f T cx)
-
-fun new_fun func_prefix t ss (Tidx, typs, dtyps, idx, terms) =
-  let
-    val f = string_of_index func_prefix idx
-    val terms' = Termtab.update (revert_types t, (f, ss)) terms
-  in (f, (Tidx, typs, dtyps, idx+1, terms')) end
-
-fun fresh_fun func_prefix t ss (cx as (_, _, _, _, terms)) =
-  (case Termtab.lookup terms (revert_types t) of
-    SOME (f, _) => (f, cx)
-  | NONE => new_fun func_prefix t ss cx)
-
-fun mk_type (_, Tfs) (d as Datatype.DtTFree _) = the (AList.lookup (op =) Tfs d)
-  | mk_type Ts (Datatype.DtType (n, ds)) = Type (n, map (mk_type Ts) ds)
-  | mk_type (Tds, _) (Datatype.DtRec i) = nth Tds i
-
-fun mk_selector ctxt Ts T n (i, d) =
-  (case Datatype_Selectors.lookup_selector ctxt (n, i+1) of
-    NONE => raise Fail ("missing selector for datatype constructor " ^ quote n)
-  | SOME m => mk_type Ts d |> (fn U => (Const (m, T --> U), U)))
-
-fun mk_constructor ctxt Ts T (n, args) =
-  let val (sels, Us) = split_list (map_index (mk_selector ctxt Ts T n) args)
-  in (Const (n, Us ---> T), sels) end
-
-fun lookup_datatype ctxt n Ts =
-  if member (op =) [@{type_name bool}, @{type_name nat}] n then NONE
-  else
-    Datatype.get_info (ProofContext.theory_of ctxt) n
-    |> Option.map (fn {descr, ...} =>
-         let
-           val Tds = map (fn (_, (tn, _, _)) => Type (tn, Ts))
-             (sort (int_ord o pairself fst) descr)
-           val Tfs = (case hd descr of (_, (_, tfs, _)) => tfs ~~ Ts)
-         in
-           descr |> map (fn (i, (_, _, cs)) =>
-             (nth Tds i, map (mk_constructor ctxt (Tds, Tfs) (nth Tds i)) cs))
-         end)
-
-fun relaxed irules = (([], irules), map (prop_of o snd) irules)
-
-fun with_context header f (ths, ts) =
-  let val (us, context) = fold_map f ts empty_context
-  in ((make_sign (header ts) context, us), make_recon ths context) end
-
-
-fun translate config ctxt comments =
-  let
-    val {prefixes, is_fol, header, has_datatypes, serialize} = config
-    val {sort_prefix, func_prefix} = prefixes
-
-    fun transT (T as TFree _) = fresh_typ T (new_typ sort_prefix true)
-      | transT (T as TVar _) = (fn _ => raise TYPE ("smt_translate", [T], []))
-      | transT (T as Type (n, Ts)) =
+    fun transT (T as TFree _) = add_typ T true
+      | transT (T as TVar _) = (fn _ => raise TYPE ("bad SMT type", [T], []))
+      | transT (T as Type _) =
           (case B.builtin_typ ctxt T of
             SOME n => pair n
-          | NONE => fresh_typ T (fn _ => fn cx =>
-              if not has_datatypes then new_typ sort_prefix true T cx
-              else
-                (case lookup_datatype ctxt n Ts of
-                  NONE => new_typ sort_prefix true T cx
-                | SOME dts =>
-                    let val cx' = new_dtyps dts cx 
-                    in (fst (the (lookup_typ cx' T)), cx') end)))
+          | NONE => add_typ T true)
 
-    and new_dtyps dts cx =
-      let
-        fun new_decl i t =
-          let val (Ts, T) = U.dest_funT i (Term.fastype_of t)
-          in
-            fold_map transT Ts ##>> transT T ##>>
-            new_fun func_prefix t NONE #>> swap
-          end
-        fun new_dtyp_decl (con, sels) =
-          new_decl (length sels) con ##>> fold_map (new_decl 1) sels #>>
-          (fn ((con', _), sels') => (con', map (apsnd snd) sels'))
-      in
-        cx
-        |> fold_map (new_typ sort_prefix false o fst) dts
-        ||>> fold_map (fold_map new_dtyp_decl o snd) dts
-        |-> (fn (ss, decls) => fn (Tidx, typs, dtyps, idx, terms) =>
-              (Tidx, typs, (ss ~~ decls) :: dtyps, idx, terms))
-      end
+    val unmarked_builtins = [@{const_name If}, @{const_name HOL.eq}]
 
     fun app n ts = SApp (n, ts)
 
@@ -361,37 +547,84 @@
             SOME (q, Ts, ps, w, b) =>
               fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>>
               trans b #>> (fn ((Ts', ps'), b') => SQua (q, Ts', ps', w, b'))
-          | NONE => raise TERM ("intermediate", [t]))
+          | NONE => raise TERM ("unsupported quantifier", [t]))
       | (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) =>
           transT T ##>> trans t1 ##>> trans t2 #>>
           (fn ((U, u1), u2) => SLet (U, u1, u2))
-      | (h as Const (c as (@{const_name distinct}, T)), ts) =>
-          (case B.builtin_fun ctxt c ts of
-            SOME (n, ts) => fold_map trans ts #>> app n
-          | NONE => transs h T ts)
-      | (h as Const (c as (_, T)), ts) =>
-          (case B.builtin_num ctxt t of
-            SOME n => pair (SApp (n, []))
-          | NONE =>
-              (case B.builtin_fun ctxt c ts of
-                SOME (n, ts') => fold_map trans ts' #>> app n
-              | NONE => transs h T ts))
-      | (h as Free (_, T), ts) => transs h T ts
+      | (Var ((n, _), _), ts) => fold_map trans ts #>> app n
+      | (u as Const (c as (n, T)), ts) =>
+          if member (op =) unmarked_builtins n then
+            (case B.builtin_fun ctxt c ts of
+              SOME (((m, _), _), us, _) => fold_map trans us #>> app m
+            | NONE => raise TERM ("not a built-in symbol", [t]))
+          else transs u T ts
+      | (u as Free (_, T), ts) => transs u T ts
       | (Bound i, []) => pair (SVar i)
-      | (Abs (_, _, t1 $ Bound 0), []) =>
-        if not (loose_bvar1 (t1, 0)) then trans t1 (* eta-reduce on the fly *)
-        else raise TERM ("smt_translate", [t])
-      | _ => raise TERM ("smt_translate", [t]))
-
+      | _ => raise TERM ("bad SMT term", [t]))
+ 
     and transs t T ts =
       let val (Us, U) = U.dest_funT (length ts) T
       in
         fold_map transT Us ##>> transT U #-> (fn Up =>
-        fresh_fun func_prefix t (SOME Up) ##>> fold_map trans ts #>> SApp)
+        add_fun t (SOME Up) ##>> fold_map trans ts #>> SApp)
       end
+
+    val (us, trx') = fold_map trans ts trx
+  in ((sign_of (header ts) dtyps trx', us), trx') end
+
+
+
+(* translation *)
+
+structure Configs = Generic_Data
+(
+  type T = (Proof.context -> config) U.dict
+  val empty = []
+  val extend = I
+  val merge = U.dict_merge fst
+)
+
+fun add_config (cs, cfg) = Configs.map (U.dict_update (cs, cfg))
+
+fun translate ctxt comments ithms =
+  let
+    val cs = SMT_Config.solver_class_of ctxt
+    val {prefixes, is_fol, header, has_datatypes, serialize} =
+      (case U.dict_get (Configs.get (Context.Proof ctxt)) cs of
+        SOME cfg => cfg ctxt
+      | NONE => error ("SMT: no translation configuration found " ^
+          "for solver class " ^ quote (U.string_of_class cs)))
+      
+    val with_datatypes =
+      has_datatypes andalso Config.get ctxt SMT_Config.datatypes
+
+    fun no_dtyps (tr_context, ctxt) ts = (([], tr_context, ctxt), ts)
+
+    val (builtins, ts1) =
+      ithms
+      |> map (HOLogic.dest_Trueprop o Thm.prop_of o snd)
+      |> map (Envir.eta_contract o Envir.beta_norm)
+      |> mark_builtins ctxt
+
+    val ((dtyps, tr_context, ctxt1), ts2) =
+      ((make_tr_context prefixes, ctxt), ts1)
+      |-> (if with_datatypes then collect_datatypes_and_records else no_dtyps)
+
+    val (ctxt2, ts3) =
+      ts2
+      |> eta_expand
+      |> lift_lambdas ctxt1
+      ||> intro_explicit_application
+
+    val ((rewrite_rules, extra_thms), ts4) =
+      (if is_fol then folify ctxt2 builtins else pair ([], [])) ts3
+
+    val rewrite_rules' = fun_app_eq :: rewrite_rules
   in
-    (if is_fol then folify ctxt else relaxed) #>
-    with_context (header ctxt) trans #>> uncurry (serialize comments)
+    (ts4, tr_context)
+    |-> intermediate header dtyps ctxt2
+    |>> uncurry (serialize comments)
+    ||> recon_of ctxt2 rewrite_rules' extra_thms ithms revert_typ revert_types
   end
 
 end
--- a/src/HOL/Tools/SMT/smt_utils.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_utils.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -13,11 +13,13 @@
   (*class dictionaries*)
   type class = string list
   val basicC: class
+  val string_of_class: class -> string
   type 'a dict = (class * 'a) Ord_List.T
   val dict_map_default: class * 'a -> ('a -> 'a) -> 'a dict -> 'a dict
   val dict_update: class * 'a -> 'a dict -> 'a dict
   val dict_merge: ('a * 'a -> 'a) -> 'a dict * 'a dict -> 'a dict
   val dict_lookup: 'a dict -> class -> 'a list
+  val dict_get: 'a dict -> class -> 'a option
 
   (*types*)
   val dest_funT: int -> typ -> typ list * typ
@@ -76,6 +78,9 @@
 
 val basicC = []
 
+fun string_of_class [] = "basic"
+  | string_of_class cs = "basic." ^ space_implode "." cs
+
 type 'a dict = (class * 'a) Ord_List.T
 
 fun class_ord ((cs1, _), (cs2, _)) = list_ord fast_string_ord (cs1, cs2)
@@ -95,6 +100,11 @@
   let fun match (cs', x) = if is_prefix (op =) cs' cs then SOME x else NONE
   in map_filter match d end
 
+fun dict_get d cs =
+  (case AList.lookup (op =) d cs of
+    NONE => (case cs of [] => NONE | _ => dict_get d (take (length cs - 1) cs))
+  | SOME x => SOME x)
+
 
 (* types *)
 
--- a/src/HOL/Tools/SMT/smtlib_interface.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Tools/SMT/smtlib_interface.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -9,7 +9,7 @@
   val smtlibC: SMT_Utils.class
   val add_logic: int * (term list -> string option) -> Context.generic ->
     Context.generic
-  val interface: SMT_Solver.interface
+  val translate_config: Proof.context -> SMT_Translate.config
   val setup: theory -> theory
 end
 
@@ -28,9 +28,13 @@
 local
   fun int_num _ i = SOME (string_of_int i)
 
-  fun distinct _ _ [t] =
+  fun distinct _ (Type (_, [Type (_, [T]), _])) [t] =
         (case try HOLogic.dest_list t of
-          SOME (ts as _ :: _) => SOME ("distinct", ts)
+          SOME (ts as _ :: _) =>
+            let
+              fun mkT U = replicate (length ts) U ---> @{typ bool}
+              val U = mkT T and U' = mkT (TVar (("'a", 0), @{sort type}))
+            in SOME ((("distinct", length ts), U), ts, U') end
         | _ => NONE)
     | distinct _ _ _ = NONE
 in
@@ -40,7 +44,6 @@
   fold (B.add_builtin_fun' smtlibC) [
     (@{const True}, "true"),
     (@{const False}, "false"),
- (* FIXME: we do not test if these constants are fully applied! *)
     (@{const Not}, "not"),
     (@{const HOL.conj}, "and"),
     (@{const HOL.disj}, "or"),
@@ -145,17 +148,17 @@
 
 (* interface *)
 
-val interface = {
-  class = smtlibC,
-  translate = {
-    prefixes = {
-      sort_prefix = "S",
-      func_prefix = "f"},
-    header = choose_logic,
-    is_fol = true,
-    has_datatypes = false,
-    serialize = serialize}}
+fun translate_config ctxt = {
+  prefixes = {
+    sort_prefix = "S",
+    func_prefix = "f"},
+  header = choose_logic ctxt,
+  is_fol = true,
+  has_datatypes = false,
+  serialize = serialize}
 
-val setup = Context.theory_map setup_builtins
+val setup = Context.theory_map (
+  setup_builtins #>
+  T.add_config (smtlibC, translate_config))
 
 end
--- a/src/HOL/Tools/SMT/z3_interface.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Tools/SMT/z3_interface.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -7,7 +7,6 @@
 signature Z3_INTERFACE =
 sig
   val smtlib_z3C: SMT_Utils.class
-  val interface: SMT_Solver.interface
   val setup: theory -> theory
 
   datatype sym = Sym of string * sym list
@@ -36,8 +35,14 @@
 (* interface *)
 
 local
-  val {translate, extra_norm, ...} = SMTLIB_Interface.interface
-  val {prefixes, is_fol, header, serialize, ...} = translate
+  fun translate_config ctxt =
+    let
+      val {prefixes, header, is_fol, serialize, ...} =
+        SMTLIB_Interface.translate_config ctxt
+    in
+      {prefixes=prefixes, header=header, is_fol=is_fol, serialize=serialize,
+        has_datatypes=true}
+    end
 
   fun is_int_div_mod @{const div (int)} = true
     | is_int_div_mod @{const mod (int)} = true
@@ -56,18 +61,10 @@
     B.add_builtin_fun' smtlib_z3C (@{const z3mod}, "mod")
 in
 
-val interface = {
-  class = smtlib_z3C,
-  translate = {
-    prefixes = prefixes,
-    is_fol = is_fol,
-    header = header,
-    has_datatypes = true,
-    serialize = serialize}}
-
 val setup = Context.theory_map (
   setup_builtins #>
-  SMT_Normalize.add_extra_norm (smtlib_z3C, add_div_mod))
+  SMT_Normalize.add_extra_norm (smtlib_z3C, add_div_mod) #>
+  SMT_Translate.add_config (smtlib_z3C, translate_config))
 
 end
 
@@ -154,11 +151,8 @@
   | mk_distinct (cts as (ct :: _)) =
       Thm.capply (U.instT' ct distinct) (mk_list (Thm.ctyp_of_term ct) cts)
 
-val access = U.mk_const_pat @{theory} @{const_name fun_app}
-  (Thm.dest_ctyp o U.destT1)
-fun mk_access array index =
-  let val cTs = Thm.dest_ctyp (Thm.ctyp_of_term array)
-  in Thm.mk_binop (U.instTs cTs access) array index end
+val access = U.mk_const_pat @{theory} @{const_name fun_app} U.destT1
+fun mk_access array = Thm.capply (U.instT' array access) array
 
 val update = U.mk_const_pat @{theory} @{const_name fun_upd}
   (Thm.dest_ctyp o U.destT1)
@@ -189,7 +183,7 @@
   | (Sym ("ite", _), [ct1, ct2, ct3]) => SOME (mk_if ct1 ct2 ct3)
   | (Sym ("=", _), [ct, cu]) => SOME (mk_eq ct cu)
   | (Sym ("distinct", _), _) => SOME (mk_distinct cts)
-  | (Sym ("select", _), [ca, ck]) => SOME (mk_access ca ck)
+  | (Sym ("select", _), [ca, ck]) => SOME (Thm.capply (mk_access ca) ck)
   | (Sym ("store", _), [ca, ck, cv]) => SOME (mk_update ca ck cv)
   | _ =>
     (case (sym, try (#T o Thm.rep_cterm o hd) cts, cts) of
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -8,7 +8,7 @@
 sig
   val add_z3_rule: thm -> Context.generic -> Context.generic
   val reconstruct: Proof.context -> SMT_Translate.recon -> string list ->
-    (int list * thm) * Proof.context
+    int list * thm
   val setup: theory -> theory
 end
 
@@ -160,15 +160,17 @@
     | _ => z3_exn ("not asserted: " ^
         quote (Syntax.string_of_term ctxt (Thm.term_of ct))))
 in
-fun prepare_assms ctxt unfolds assms =
+fun prepare_assms ctxt rewrite_rules assms =
   let
-    val unfolds' = rewrites I ctxt [L.rewrite_true] unfolds
+    val eqs = rewrites I ctxt [L.rewrite_true] rewrite_rules
     val assms' =
-      rewrites apsnd ctxt (union Thm.eq_thm unfolds' prep_rules) assms
-  in (unfolds', T.thm_net_of snd assms') end
+      assms
+      |> rewrites apsnd ctxt (union Thm.eq_thm eqs prep_rules)
+      |> map (apsnd (Conv.fconv_rule Thm.eta_conversion))
+  in (eqs, T.thm_net_of snd assms') end
 
-fun asserted ctxt (unfolds, assms) ct =
-  let val revert_conv = rewrite_conv ctxt unfolds
+fun asserted ctxt (eqs, assms) ct =
+  let val revert_conv = rewrite_conv ctxt eqs then_conv Thm.eta_conversion
   in Thm (T.with_conv revert_conv (snd o lookup_assm ctxt assms) ct) end
 
 fun find_assm ctxt (unfolds, assms) ct =
@@ -835,6 +837,22 @@
       (fn idx => result o lookup idx o pair ctxt o Inttab.map (K Unproved))
     end
 
+  val disch_rules = map (pair false)
+    [@{thm allI}, @{thm refl}, @{thm reflexive}]
+
+  fun disch_assm thm =
+    if Thm.nprems_of thm = 0 then Drule.flexflex_unique thm
+    else
+      (case Seq.pull (Thm.biresolution false disch_rules 1 thm) of
+        SOME (thm', _) => disch_assm thm'
+      | NONE => raise THM ("failed to discharge premise", 1, [thm]))
+
+  fun discharge outer_ctxt (thm, inner_ctxt) =
+    thm
+    |> singleton (ProofContext.export inner_ctxt outer_ctxt)
+    |> tap (tracing o prefix "final goal: " o PolyML.makestring)
+    |> disch_assm    
+
   fun filter_assms ctxt assms ptab =
     let
       fun step r ct =
@@ -851,14 +869,18 @@
     in lookup end
 in
 
-fun reconstruct ctxt {typs, terms, unfolds, assms} output =
+fun reconstruct outer_ctxt recon output =
   let
-    val (idx, (ptab, vars, cx)) = P.parse ctxt typs terms output
-    val assms' = prepare_assms cx unfolds assms
+    val {context=ctxt, typs, terms, rewrite_rules, assms} = recon
+    val (idx, (ptab, vars, ctxt')) = P.parse ctxt typs terms output
+    val assms' = prepare_assms ctxt' rewrite_rules assms
   in
-    if Config.get cx SMT_Config.filter_only_facts
-    then ((filter_assms cx assms' ptab idx [], @{thm TrueI}), cx)
-    else apfst (pair []) (prove cx assms' vars idx ptab)
+    if Config.get ctxt' SMT_Config.filter_only_facts then
+      (filter_assms ctxt' assms' ptab idx [], @{thm TrueI})
+    else
+      prove ctxt' assms' vars idx ptab
+      |> discharge outer_ctxt
+      |> pair []
   end
 
 end
--- a/src/HOL/Word/Tools/smt_word.ML	Wed Dec 15 08:39:24 2010 +0100
+++ b/src/HOL/Word/Tools/smt_word.ML	Wed Dec 15 10:12:44 2010 +0100
@@ -59,12 +59,19 @@
     | word_num _ _ = NONE
 
   fun if_fixed n T ts =
-    let val (Ts, T) = Term.strip_type T
-    in if forall (can dest_wordT) (T :: Ts) then SOME (n, ts) else NONE end
+    let val (Us, U) = Term.strip_type T
+    in
+      if forall (can dest_wordT) (U :: Us) then
+        SOME (((n, length Us), T), ts, T)
+      else NONE
+    end
 
   fun if_fixed' n T ts =
-    if forall (can dest_wordT) (Term.binder_types T) then SOME (n, ts)
-    else NONE
+    let val Ts = Term.binder_types T
+    in
+      if forall (can dest_wordT) Ts then SOME (((n, length Ts), T), ts, T)
+      else NONE
+    end
 
   fun add_word_fun f (t, n) =
     B.add_builtin_fun smtlibC (Term.dest_Const t, K (f n))
@@ -79,28 +86,37 @@
     (dest_word_funT (Term.range_type T), dest_nat ts)
 
   fun shift n T ts =
-    let val U = Term.domain_type T
+    let
+      val U = Term.domain_type T
+      val T' = [U, U] ---> U
     in
-      (case (can dest_wordT U, ts) of
+      (case (can dest_wordT T', ts) of
         (true, [t, u]) =>
           (case try HOLogic.dest_number u of
-            SOME (_,i) => SOME (n, [t, HOLogic.mk_number U i])
+            SOME (_, i) => SOME (((n, 2), T'), [t, HOLogic.mk_number T' i], T')
           | NONE => NONE)  (* FIXME: also support non-numerical shifts *)
       | _ => NONE)
     end
 
   fun extract n T ts =
     try dest_nat_word_funT (T, ts)
-    |> Option.map (fn ((_, i), (lb, ts')) => (index2 n (i + lb - 1) lb, ts'))
+    |> Option.map (fn ((_, i), (lb, ts')) =>
+         let val T' = Term.range_type T
+         in (((index2 n (i + lb - 1) lb, 1), T'), ts', T') end)
 
   fun extend n T ts =
     (case try dest_word_funT T of
-      SOME (i, j) => if j-i >= 0 then SOME (index1 n (j-i), ts) else NONE
+      SOME (i, j) =>
+        if j-i >= 0 then SOME (((index1 n (j-i), 1), T), ts, T)
+        else NONE
     | _ => NONE)
 
   fun rotate n T ts =
-    try dest_nat ts
-    |> Option.map (fn (i, ts') => (index1 n i, ts'))
+    let val T' = Term.range_type T
+    in
+      try dest_nat ts
+      |> Option.map (fn (i, ts') => (((index1 n i, 1), T'), ts', T'))
+    end
 in
 
 val setup_builtins =