src/HOL/Tools/SMT/smt_translate.ML
changeset 41127 2ea84c8535c6
parent 41123 3bb9be510a9d
child 41165 ceb81a08534e
--- 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