src/HOL/Tools/SMT/smt_translate.ML
author boehmes
Wed, 20 Jul 2011 15:09:53 +0200
changeset 44803 b2218b8265ea
parent 44800 61d432e51aff
child 47400 a018d542e0ae
permissions -rw-r--r--
removed debugging facilities accidentally left in the committed code
     1 (*  Title:      HOL/Tools/SMT/smt_translate.ML
     2     Author:     Sascha Boehme, TU Muenchen
     3 
     4 Translate theorems into an SMT intermediate format and serialize them.
     5 *)
     6 
     7 signature SMT_TRANSLATE =
     8 sig
     9   (*intermediate term structure*)
    10   datatype squant = SForall | SExists
    11   datatype 'a spattern = SPat of 'a list | SNoPat of 'a list
    12   datatype sterm =
    13     SVar of int |
    14     SApp of string * sterm list |
    15     SLet of string * sterm * sterm |
    16     SQua of squant * string list * sterm spattern list * int option * sterm
    17 
    18   (*translation configuration*)
    19   type prefixes = {sort_prefix: string, func_prefix: string}
    20   type sign = {
    21     header: string list,
    22     sorts: string list,
    23     dtyps: (string * (string * (string * string) list) list) list list,
    24     funcs: (string * (string list * string)) list }
    25   type config = {
    26     prefixes: prefixes,
    27     header: term list -> string list,
    28     is_fol: bool,
    29     has_datatypes: bool,
    30     serialize: string list -> sign -> sterm list -> string }
    31   type recon = {
    32     context: Proof.context,
    33     typs: typ Symtab.table,
    34     terms: term Symtab.table,
    35     rewrite_rules: thm list,
    36     assms: (int * thm) list }
    37 
    38   (*translation*)
    39   val add_config: SMT_Utils.class * (Proof.context -> config) ->
    40     Context.generic -> Context.generic 
    41   val translate: Proof.context -> string list -> (int * thm) list ->
    42     string * recon
    43 end
    44 
    45 structure SMT_Translate: SMT_TRANSLATE =
    46 struct
    47 
    48 
    49 (* intermediate term structure *)
    50 
    51 datatype squant = SForall | SExists
    52 
    53 datatype 'a spattern = SPat of 'a list | SNoPat of 'a list
    54 
    55 datatype sterm =
    56   SVar of int |
    57   SApp of string * sterm list |
    58   SLet of string * sterm * sterm |
    59   SQua of squant * string list * sterm spattern list * int option * sterm
    60 
    61 
    62 
    63 (* translation configuration *)
    64 
    65 type prefixes = {sort_prefix: string, func_prefix: string}
    66 
    67 type sign = {
    68   header: string list,
    69   sorts: string list,
    70   dtyps: (string * (string * (string * string) list) list) list list,
    71   funcs: (string * (string list * string)) list }
    72 
    73 type config = {
    74   prefixes: prefixes,
    75   header: term list -> string list,
    76   is_fol: bool,
    77   has_datatypes: bool,
    78   serialize: string list -> sign -> sterm list -> string }
    79 
    80 type recon = {
    81   context: Proof.context,
    82   typs: typ Symtab.table,
    83   terms: term Symtab.table,
    84   rewrite_rules: thm list,
    85   assms: (int * thm) list }
    86 
    87 
    88 
    89 (* translation context *)
    90 
    91 fun make_tr_context {sort_prefix, func_prefix} =
    92   (sort_prefix, 1, Typtab.empty, func_prefix, 1, Termtab.empty)
    93 
    94 fun string_of_index pre i = pre ^ string_of_int i
    95 
    96 fun add_typ T proper (cx as (sp, Tidx, typs, fp, idx, terms)) =
    97   (case Typtab.lookup typs T of
    98     SOME (n, _) => (n, cx)
    99   | NONE =>
   100       let
   101         val n = string_of_index sp Tidx
   102         val typs' = Typtab.update (T, (n, proper)) typs
   103       in (n, (sp, Tidx+1, typs', fp, idx, terms)) end)
   104 
   105 fun add_fun t sort (cx as (sp, Tidx, typs, fp, idx, terms)) =
   106   (case Termtab.lookup terms t of
   107     SOME (n, _) => (n, cx)
   108   | NONE => 
   109       let
   110         val n = string_of_index fp idx
   111         val terms' = Termtab.update (t, (n, sort)) terms
   112       in (n, (sp, Tidx, typs, fp, idx+1, terms')) end)
   113 
   114 fun sign_of header dtyps (_, _, typs, _, _, terms) = {
   115   header = header,
   116   sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [],
   117   dtyps = dtyps,
   118   funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms []}
   119 
   120 fun recon_of ctxt rules thms ithms (_, _, typs, _, _, terms) =
   121   let
   122     fun add_typ (T, (n, _)) = Symtab.update (n, T)
   123     val typs' = Typtab.fold add_typ typs Symtab.empty
   124 
   125     fun add_fun (t, (n, _)) = Symtab.update (n, t)
   126     val terms' = Termtab.fold add_fun terms Symtab.empty
   127 
   128     val assms = map (pair ~1) thms @ ithms
   129   in
   130     {context=ctxt, typs=typs', terms=terms', rewrite_rules=rules, assms=assms}
   131   end
   132 
   133 
   134 
   135 (* preprocessing *)
   136 
   137 (** datatype declarations **)
   138 
   139 fun collect_datatypes_and_records (tr_context, ctxt) ts =
   140   let
   141     val (declss, ctxt') =
   142       fold (Term.fold_types SMT_Datatypes.add_decls) ts ([], ctxt)
   143 
   144     fun is_decl_typ T = exists (exists (equal T o fst)) declss
   145 
   146     fun add_typ' T proper =
   147       (case SMT_Builtin.dest_builtin_typ ctxt' T of
   148         SOME n => pair n
   149       | NONE => add_typ T proper)
   150 
   151     fun tr_select sel =
   152       let val T = Term.range_type (Term.fastype_of sel)
   153       in add_fun sel NONE ##>> add_typ' T (not (is_decl_typ T)) end
   154     fun tr_constr (constr, selects) =
   155       add_fun constr NONE ##>> fold_map tr_select selects
   156     fun tr_typ (T, cases) = add_typ' T false ##>> fold_map tr_constr cases
   157     val (declss', tr_context') = fold_map (fold_map tr_typ) declss tr_context
   158 
   159     fun add (constr, selects) =
   160       Termtab.update (constr, length selects) #>
   161       fold (Termtab.update o rpair 1) selects
   162     val funcs = fold (fold (fold add o snd)) declss Termtab.empty
   163   in ((funcs, declss', tr_context', ctxt'), ts) end
   164     (* FIXME: also return necessary datatype and record theorems *)
   165 
   166 
   167 (** eta-expand quantifiers, let expressions and built-ins *)
   168 
   169 local
   170   fun eta f T t = Abs (Name.uu, T, f (Term.incr_boundvars 1 t $ Bound 0))
   171 
   172   fun exp f T = eta f (Term.domain_type (Term.domain_type T))
   173 
   174   fun exp2 T q =
   175     let val U = Term.domain_type T
   176     in Abs (Name.uu, U, q $ eta I (Term.domain_type U) (Bound 0)) end
   177 
   178   fun exp2' T l =
   179     let val (U1, U2) = Term.dest_funT T ||> Term.domain_type
   180     in Abs (Name.uu, U1, eta I U2 (l $ Bound 0)) end
   181 
   182   fun expf k i T t =
   183     let val Ts = drop i (fst (SMT_Utils.dest_funT k T))
   184     in
   185       Term.incr_boundvars (length Ts) t
   186       |> fold_rev (fn i => fn u => u $ Bound i) (0 upto length Ts - 1)
   187       |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts
   188     end
   189 in
   190 
   191 fun eta_expand ctxt is_fol funcs =
   192   let
   193     fun exp_func t T ts =
   194       (case Termtab.lookup funcs t of
   195         SOME k =>
   196           Term.list_comb (t, ts)
   197           |> k <> length ts ? expf k (length ts) T
   198       | NONE => Term.list_comb (t, ts))
   199 
   200     fun expand ((q as Const (@{const_name All}, _)) $ Abs a) = q $ abs_expand a
   201       | expand ((q as Const (@{const_name All}, T)) $ t) = q $ exp expand T t
   202       | expand (q as Const (@{const_name All}, T)) = exp2 T q
   203       | expand ((q as Const (@{const_name Ex}, _)) $ Abs a) = q $ abs_expand a
   204       | expand ((q as Const (@{const_name Ex}, T)) $ t) = q $ exp expand T t
   205       | expand (q as Const (@{const_name Ex}, T)) = exp2 T q
   206       | expand ((l as Const (@{const_name Let}, _)) $ t $ Abs a) =
   207           if is_fol then expand (Term.betapply (Abs a, t))
   208           else l $ expand t $ abs_expand a
   209       | expand ((l as Const (@{const_name Let}, T)) $ t $ u) =
   210           if is_fol then expand (u $ t)
   211           else l $ expand t $ exp expand (Term.range_type T) u
   212       | expand ((l as Const (@{const_name Let}, T)) $ t) =
   213           if is_fol then
   214             let val U = Term.domain_type (Term.range_type T)
   215             in Abs (Name.uu, U, Bound 0 $ Term.incr_boundvars 1 t) end
   216           else exp2 T (l $ expand t)
   217       | expand (l as Const (@{const_name Let}, T)) =
   218           if is_fol then 
   219             let val U = Term.domain_type (Term.range_type T)
   220             in
   221               Abs (Name.uu, Term.domain_type T, Abs (Name.uu, U,
   222                 Bound 0 $ Bound 1))
   223             end
   224           else exp2' T l
   225       | expand t =
   226           (case Term.strip_comb t of
   227             (u as Const (c as (_, T)), ts) =>
   228               (case SMT_Builtin.dest_builtin ctxt c ts of
   229                 SOME (_, k, us, mk) =>
   230                   if k = length us then mk (map expand us)
   231                   else expf k (length ts) T (mk (map expand us))
   232               | NONE => exp_func u T (map expand ts))
   233           | (u as Free (_, T), ts) => exp_func u T (map expand ts)
   234           | (Abs a, ts) => Term.list_comb (abs_expand a, map expand ts)
   235           | (u, ts) => Term.list_comb (u, map expand ts))
   236 
   237     and abs_expand (n, T, t) = Abs (n, T, expand t)
   238   
   239   in map expand end
   240 
   241 end
   242 
   243 
   244 (** introduce explicit applications **)
   245 
   246 local
   247   (*
   248     Make application explicit for functions with varying number of arguments.
   249   *)
   250 
   251   fun add t i = apfst (Termtab.map_default (t, i) (Integer.min i))
   252   fun add_type T = apsnd (Typtab.update (T, ()))
   253 
   254   fun min_arities t =
   255     (case Term.strip_comb t of
   256       (u as Const _, ts) => add u (length ts) #> fold min_arities ts
   257     | (u as Free _, ts) => add u (length ts) #> fold min_arities ts
   258     | (Abs (_, T, u), ts) => add_type T #> min_arities u #> fold min_arities ts
   259     | (_, ts) => fold min_arities ts)
   260 
   261   fun minimize types t i =
   262     let
   263       fun find_min j [] _ = j
   264         | find_min j (U :: Us) T =
   265             if Typtab.defined types T then j
   266             else find_min (j + 1) Us (U --> T)
   267 
   268       val (Ts, T) = Term.strip_type (Term.type_of t)
   269     in find_min 0 (take i (rev Ts)) T end
   270 
   271   fun app u (t, T) =
   272     (Const (@{const_name SMT.fun_app}, T --> T) $ t $ u, Term.range_type T)
   273 
   274   fun apply i t T ts =
   275     let
   276       val (ts1, ts2) = chop i ts
   277       val (_, U) = SMT_Utils.dest_funT i T
   278     in fst (fold app ts2 (Term.list_comb (t, ts1), U)) end
   279 in
   280 
   281 fun intro_explicit_application ctxt funcs ts =
   282   let
   283     val (arities, types) = fold min_arities ts (Termtab.empty, Typtab.empty)
   284     val arities' = Termtab.map (minimize types) arities
   285 
   286     fun app_func t T ts =
   287       if is_some (Termtab.lookup funcs t) then Term.list_comb (t, ts)
   288       else apply (the (Termtab.lookup arities' t)) t T ts
   289 
   290     fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t))
   291 
   292     fun traverse Ts t =
   293       (case Term.strip_comb t of
   294         (q as Const (@{const_name All}, _), [Abs (x, T, u)]) =>
   295           q $ Abs (x, T, in_trigger (T :: Ts) u)
   296       | (q as Const (@{const_name Ex}, _), [Abs (x, T, u)]) =>
   297           q $ Abs (x, T, in_trigger (T :: Ts) u)
   298       | (q as Const (@{const_name Let}, _), [u1, u2 as Abs _]) =>
   299           q $ traverse Ts u1 $ traverse Ts u2
   300       | (u as Const (c as (_, T)), ts) =>
   301           (case SMT_Builtin.dest_builtin ctxt c ts of
   302             SOME (_, _, us, mk) => mk (map (traverse Ts) us)
   303           | NONE => app_func u T (map (traverse Ts) ts))
   304       | (u as Free (_, T), ts) => app_func u T (map (traverse Ts) ts)
   305       | (u as Bound i, ts) => apply 0 u (nth Ts i) (map (traverse Ts) ts)
   306       | (Abs (n, T, u), ts) => traverses Ts (Abs (n, T, traverse (T::Ts) u)) ts
   307       | (u, ts) => traverses Ts u ts)
   308     and in_trigger Ts ((c as @{const SMT.trigger}) $ p $ t) =
   309           c $ in_pats Ts p $ in_weight Ts t
   310       | in_trigger Ts t = in_weight Ts t
   311     and in_pats Ts ps =
   312       in_list @{typ "SMT.pattern list"}
   313         (in_list @{typ SMT.pattern} (in_pat Ts)) ps
   314     and in_pat Ts ((p as Const (@{const_name SMT.pat}, _)) $ t) =
   315           p $ traverse Ts t
   316       | in_pat Ts ((p as Const (@{const_name SMT.nopat}, _)) $ t) =
   317           p $ traverse Ts t
   318       | in_pat _ t = raise TERM ("bad pattern", [t])
   319     and in_weight Ts ((c as @{const SMT.weight}) $ w $ t) =
   320           c $ w $ traverse Ts t
   321       | in_weight Ts t = traverse Ts t 
   322     and traverses Ts t ts = Term.list_comb (t, map (traverse Ts) ts)
   323   in map (traverse []) ts end
   324 
   325 val fun_app_eq = mk_meta_eq @{thm SMT.fun_app_def}
   326 
   327 end
   328 
   329 
   330 (** map HOL formulas to FOL formulas (i.e., separate formulas froms terms) **)
   331 
   332 local
   333   val term_bool = @{lemma "SMT.term_true ~= SMT.term_false"
   334     by (simp add: SMT.term_true_def SMT.term_false_def)}
   335 
   336   val is_quant = member (op =) [@{const_name All}, @{const_name Ex}]
   337 
   338   val fol_rules = [
   339     Let_def,
   340     mk_meta_eq @{thm SMT.term_true_def},
   341     mk_meta_eq @{thm SMT.term_false_def},
   342     @{lemma "P = True == P" by (rule eq_reflection) simp},
   343     @{lemma "if P then True else False == P" by (rule eq_reflection) simp}]
   344 
   345   fun as_term t = @{const HOL.eq (bool)} $ t $ @{const SMT.term_true}
   346 
   347   fun wrap_in_if t =
   348     @{const If (bool)} $ t $ @{const SMT.term_true} $ @{const SMT.term_false}
   349 
   350   fun is_builtin_conn_or_pred ctxt c ts =
   351     is_some (SMT_Builtin.dest_builtin_conn ctxt c ts) orelse
   352     is_some (SMT_Builtin.dest_builtin_pred ctxt c ts)
   353 
   354   fun builtin b ctxt c ts =
   355     (case (Const c, ts) of
   356       (@{const HOL.eq (bool)}, [t, u]) =>
   357         if t = @{const SMT.term_true} orelse u = @{const SMT.term_true} then
   358           SMT_Builtin.dest_builtin_eq ctxt t u
   359         else b ctxt c ts
   360     | _ => b ctxt c ts)
   361 in
   362 
   363 fun folify ctxt =
   364   let
   365     fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t))
   366 
   367     fun in_term t =
   368       (case Term.strip_comb t of
   369         (@{const True}, []) => @{const SMT.term_true}
   370       | (@{const False}, []) => @{const SMT.term_false}
   371       | (u as Const (@{const_name If}, _), [t1, t2, t3]) =>
   372           u $ in_form t1 $ in_term t2 $ in_term t3
   373       | (Const (c as (n, _)), ts) =>
   374           if is_builtin_conn_or_pred ctxt c ts then wrap_in_if (in_form t)
   375           else  if is_quant n then wrap_in_if (in_form t)
   376           else Term.list_comb (Const c, map in_term ts)
   377       | (Free c, ts) => Term.list_comb (Free c, map in_term ts)
   378       | _ => t)
   379 
   380     and in_weight ((c as @{const SMT.weight}) $ w $ t) = c $ w $ in_form t
   381       | in_weight t = in_form t 
   382 
   383     and in_pat ((p as Const (@{const_name SMT.pat}, _)) $ t) = p $ in_term t
   384       | in_pat ((p as Const (@{const_name SMT.nopat}, _)) $ t) = p $ in_term t
   385       | in_pat t = raise TERM ("bad pattern", [t])
   386 
   387     and in_pats ps =
   388       in_list @{typ "SMT.pattern list"} (in_list @{typ SMT.pattern} in_pat) ps
   389 
   390     and in_trigger ((c as @{const SMT.trigger}) $ p $ t) =
   391           c $ in_pats p $ in_weight t
   392       | in_trigger t = in_weight t
   393 
   394     and in_form t =
   395       (case Term.strip_comb t of
   396         (q as Const (qn, _), [Abs (n, T, u)]) =>
   397           if is_quant qn then q $ Abs (n, T, in_trigger u)
   398           else as_term (in_term t)
   399       | (Const c, ts) =>
   400           (case SMT_Builtin.dest_builtin_conn ctxt c ts of
   401             SOME (_, _, us, mk) => mk (map in_form us)
   402           | NONE =>
   403               (case SMT_Builtin.dest_builtin_pred ctxt c ts of
   404                 SOME (_, _, us, mk) => mk (map in_term us)
   405               | NONE => as_term (in_term t)))
   406       | _ => as_term (in_term t))
   407   in
   408     map in_form #>
   409     cons (SMT_Utils.prop_of term_bool) #>
   410     pair (fol_rules, [term_bool], builtin)
   411   end
   412 
   413 end
   414 
   415 
   416 (* translation into intermediate format *)
   417 
   418 (** utility functions **)
   419 
   420 val quantifier = (fn
   421     @{const_name All} => SOME SForall
   422   | @{const_name Ex} => SOME SExists
   423   | _ => NONE)
   424 
   425 fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) =
   426       if q = qname then group_quant qname (T :: Ts) u else (Ts, t)
   427   | group_quant _ Ts t = (Ts, t)
   428 
   429 fun dest_weight (@{const SMT.weight} $ w $ t) =
   430       (SOME (snd (HOLogic.dest_number w)), t)
   431   | dest_weight t = (NONE, t)
   432 
   433 fun dest_pat (Const (@{const_name SMT.pat}, _) $ t) = (t, true)
   434   | dest_pat (Const (@{const_name SMT.nopat}, _) $ t) = (t, false)
   435   | dest_pat t = raise TERM ("bad pattern", [t])
   436 
   437 fun dest_pats [] = I
   438   | dest_pats ts =
   439       (case map dest_pat ts |> split_list ||> distinct (op =) of
   440         (ps, [true]) => cons (SPat ps)
   441       | (ps, [false]) => cons (SNoPat ps)
   442       | _ => raise TERM ("bad multi-pattern", ts))
   443 
   444 fun dest_trigger (@{const SMT.trigger} $ tl $ t) =
   445       (rev (fold (dest_pats o HOLogic.dest_list) (HOLogic.dest_list tl) []), t)
   446   | dest_trigger t = ([], t)
   447 
   448 fun dest_quant qn T t = quantifier qn |> Option.map (fn q =>
   449   let
   450     val (Ts, u) = group_quant qn [T] t
   451     val (ps, p) = dest_trigger u
   452     val (w, b) = dest_weight p
   453   in (q, rev Ts, ps, w, b) end)
   454 
   455 fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat
   456   | fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat
   457 
   458 
   459 (** translation from Isabelle terms into SMT intermediate terms **)
   460 
   461 fun intermediate header dtyps builtin ctxt ts trx =
   462   let
   463     fun transT (T as TFree _) = add_typ T true
   464       | transT (T as TVar _) = (fn _ => raise TYPE ("bad SMT type", [T], []))
   465       | transT (T as Type _) =
   466           (case SMT_Builtin.dest_builtin_typ ctxt T of
   467             SOME n => pair n
   468           | NONE => add_typ T true)
   469 
   470     fun app n ts = SApp (n, ts)
   471 
   472     fun trans t =
   473       (case Term.strip_comb t of
   474         (Const (qn, _), [Abs (_, T, t1)]) =>
   475           (case dest_quant qn T t1 of
   476             SOME (q, Ts, ps, w, b) =>
   477               fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>>
   478               trans b #>> (fn ((Ts', ps'), b') => SQua (q, Ts', ps', w, b'))
   479           | NONE => raise TERM ("unsupported quantifier", [t]))
   480       | (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) =>
   481           transT T ##>> trans t1 ##>> trans t2 #>>
   482           (fn ((U, u1), u2) => SLet (U, u1, u2))
   483       | (u as Const (c as (_, T)), ts) =>
   484           (case builtin ctxt c ts of
   485             SOME (n, _, us, _) => fold_map trans us #>> app n
   486           | NONE => transs u T ts)
   487       | (u as Free (_, T), ts) => transs u T ts
   488       | (Bound i, []) => pair (SVar i)
   489       | _ => raise TERM ("bad SMT term", [t]))
   490  
   491     and transs t T ts =
   492       let val (Us, U) = SMT_Utils.dest_funT (length ts) T
   493       in
   494         fold_map transT Us ##>> transT U #-> (fn Up =>
   495         add_fun t (SOME Up) ##>> fold_map trans ts #>> SApp)
   496       end
   497 
   498     val (us, trx') = fold_map trans ts trx
   499   in ((sign_of (header ts) dtyps trx', us), trx') end
   500 
   501 
   502 
   503 (* translation *)
   504 
   505 structure Configs = Generic_Data
   506 (
   507   type T = (Proof.context -> config) SMT_Utils.dict
   508   val empty = []
   509   val extend = I
   510   fun merge data = SMT_Utils.dict_merge fst data
   511 )
   512 
   513 fun add_config (cs, cfg) = Configs.map (SMT_Utils.dict_update (cs, cfg))
   514 
   515 fun get_config ctxt = 
   516   let val cs = SMT_Config.solver_class_of ctxt
   517   in
   518     (case SMT_Utils.dict_get (Configs.get (Context.Proof ctxt)) cs of
   519       SOME cfg => cfg ctxt
   520     | NONE => error ("SMT: no translation configuration found " ^
   521         "for solver class " ^ quote (SMT_Utils.string_of_class cs)))
   522   end
   523 
   524 fun translate ctxt comments ithms =
   525   let
   526     val {prefixes, is_fol, header, has_datatypes, serialize} = get_config ctxt
   527 
   528     val with_datatypes =
   529       has_datatypes andalso Config.get ctxt SMT_Config.datatypes
   530 
   531     fun no_dtyps (tr_context, ctxt) ts =
   532       ((Termtab.empty, [], tr_context, ctxt), ts)
   533 
   534     val ts1 = map (Envir.beta_eta_contract o SMT_Utils.prop_of o snd) ithms
   535 
   536     val ((funcs, dtyps, tr_context, ctxt1), ts2) =
   537       ((make_tr_context prefixes, ctxt), ts1)
   538       |-> (if with_datatypes then collect_datatypes_and_records else no_dtyps)
   539 
   540     fun is_binder (Const (@{const_name Let}, _) $ _) = true
   541       | is_binder t = Lambda_Lifting.is_quantifier t
   542 
   543     fun mk_trigger ((q as Const (@{const_name All}, _)) $ Abs (n, T, t)) =
   544           q $ Abs (n, T, mk_trigger t)
   545       | mk_trigger (eq as (Const (@{const_name HOL.eq}, T) $ lhs $ _)) =
   546           Term.domain_type T --> @{typ SMT.pattern}
   547           |> (fn T => Const (@{const_name SMT.pat}, T) $ lhs)
   548           |> HOLogic.mk_list @{typ SMT.pattern} o single
   549           |> HOLogic.mk_list @{typ "SMT.pattern list"} o single
   550           |> (fn t => @{const SMT.trigger} $ t $ eq)
   551       | mk_trigger t = t
   552 
   553     val (ctxt2, ts3) =
   554       ts2
   555       |> eta_expand ctxt1 is_fol funcs
   556       |> rpair ctxt1
   557       |-> Lambda_Lifting.lift_lambdas NONE is_binder
   558       |-> (fn (ts', defs) => fn ctxt' =>
   559           map mk_trigger defs @ ts'
   560           |> intro_explicit_application ctxt' funcs 
   561           |> pair ctxt')
   562 
   563     val ((rewrite_rules, extra_thms, builtin), ts4) =
   564       (if is_fol then folify ctxt2 else pair ([], [], I)) ts3
   565 
   566     val rewrite_rules' = fun_app_eq :: rewrite_rules
   567   in
   568     (ts4, tr_context)
   569     |-> intermediate header dtyps (builtin SMT_Builtin.dest_builtin) ctxt2
   570     |>> uncurry (serialize comments)
   571     ||> recon_of ctxt2 rewrite_rules' extra_thms ithms
   572   end
   573 
   574 end