boehmes@36898: (* Title: HOL/Tools/SMT/smt_translate.ML boehmes@36898: Author: Sascha Boehme, TU Muenchen boehmes@36898: boehmes@36898: Translate theorems into an SMT intermediate format and serialize them. boehmes@36898: *) boehmes@36898: boehmes@36898: signature SMT_TRANSLATE = boehmes@36898: sig boehmes@41123: (*intermediate term structure*) boehmes@36898: datatype squant = SForall | SExists boehmes@36898: datatype 'a spattern = SPat of 'a list | SNoPat of 'a list boehmes@36898: datatype sterm = boehmes@36898: SVar of int | boehmes@36898: SApp of string * sterm list | boehmes@36898: SLet of string * sterm * sterm | boehmes@40664: SQua of squant * string list * sterm spattern list * int option * sterm boehmes@36898: boehmes@41127: (*translation configuration*) boehmes@36898: type prefixes = {sort_prefix: string, func_prefix: string} boehmes@36898: type sign = { boehmes@36899: header: string list, boehmes@36898: sorts: string list, boehmes@39298: dtyps: (string * (string * (string * string) list) list) list list, boehmes@36898: funcs: (string * (string list * string)) list } boehmes@36898: type config = { boehmes@36898: prefixes: prefixes, boehmes@41127: header: term list -> string list, boehmes@41059: is_fol: bool, boehmes@41059: has_datatypes: bool, boehmes@36898: serialize: string list -> sign -> sterm list -> string } boehmes@36898: type recon = { boehmes@41127: context: Proof.context, boehmes@36898: typs: typ Symtab.table, boehmes@36898: terms: term Symtab.table, boehmes@41127: rewrite_rules: thm list, boehmes@40161: assms: (int * thm) list } boehmes@36898: boehmes@41127: (*translation*) boehmes@41127: val add_config: SMT_Utils.class * (Proof.context -> config) -> boehmes@41127: Context.generic -> Context.generic boehmes@41127: val translate: Proof.context -> string list -> (int * thm) list -> boehmes@36898: string * recon boehmes@36898: end boehmes@36898: boehmes@36898: structure SMT_Translate: SMT_TRANSLATE = boehmes@36898: struct boehmes@36898: boehmes@40663: structure U = SMT_Utils boehmes@41059: structure B = SMT_Builtin boehmes@40663: boehmes@40663: boehmes@36898: (* intermediate term structure *) boehmes@36898: boehmes@36898: datatype squant = SForall | SExists boehmes@36898: boehmes@36898: datatype 'a spattern = SPat of 'a list | SNoPat of 'a list boehmes@36898: boehmes@36898: datatype sterm = boehmes@36898: SVar of int | boehmes@36898: SApp of string * sterm list | boehmes@36898: SLet of string * sterm * sterm | boehmes@40664: SQua of squant * string list * sterm spattern list * int option * sterm boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41127: (* translation configuration *) boehmes@36898: boehmes@36898: type prefixes = {sort_prefix: string, func_prefix: string} boehmes@36898: boehmes@36898: type sign = { boehmes@36899: header: string list, boehmes@36898: sorts: string list, boehmes@39298: dtyps: (string * (string * (string * string) list) list) list list, boehmes@36898: funcs: (string * (string list * string)) list } boehmes@36898: boehmes@36898: type config = { boehmes@36898: prefixes: prefixes, boehmes@41127: header: term list -> string list, boehmes@41059: is_fol: bool, boehmes@41059: has_datatypes: bool, boehmes@36898: serialize: string list -> sign -> sterm list -> string } boehmes@36898: boehmes@36898: type recon = { boehmes@41127: context: Proof.context, boehmes@36898: typs: typ Symtab.table, boehmes@36898: terms: term Symtab.table, boehmes@41127: rewrite_rules: thm list, boehmes@40161: assms: (int * thm) list } boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41127: (* translation context *) boehmes@41127: boehmes@41127: fun make_tr_context {sort_prefix, func_prefix} = boehmes@41127: (sort_prefix, 1, Typtab.empty, func_prefix, 1, Termtab.empty) boehmes@41127: boehmes@41127: fun string_of_index pre i = pre ^ string_of_int i boehmes@41127: boehmes@41127: fun add_typ T proper (cx as (sp, Tidx, typs, fp, idx, terms)) = boehmes@41127: (case Typtab.lookup typs T of boehmes@41127: SOME (n, _) => (n, cx) boehmes@41127: | NONE => boehmes@41127: let boehmes@41127: val n = string_of_index sp Tidx boehmes@41127: val typs' = Typtab.update (T, (n, proper)) typs boehmes@41127: in (n, (sp, Tidx+1, typs', fp, idx, terms)) end) boehmes@41127: boehmes@41127: fun add_fun t sort (cx as (sp, Tidx, typs, fp, idx, terms)) = boehmes@41127: (case Termtab.lookup terms t of boehmes@41127: SOME (n, _) => (n, cx) boehmes@41127: | NONE => boehmes@41127: let boehmes@41127: val n = string_of_index fp idx boehmes@41127: val terms' = Termtab.update (t, (n, sort)) terms boehmes@41127: in (n, (sp, Tidx, typs, fp, idx+1, terms')) end) boehmes@41127: boehmes@41127: fun sign_of header dtyps (_, _, typs, _, _, terms) = { boehmes@41127: header = header, boehmes@41127: sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [], boehmes@41127: dtyps = dtyps, boehmes@41127: funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms []} boehmes@41127: boehmes@41127: fun recon_of ctxt rules thms ithms revertT revert (_, _, typs, _, _, terms) = boehmes@41127: let boehmes@41127: fun add_typ (T, (n, _)) = Symtab.update (n, revertT T) boehmes@41127: val typs' = Typtab.fold add_typ typs Symtab.empty boehmes@41127: boehmes@41127: fun add_fun (t, (n, _)) = Symtab.update (n, revert t) boehmes@41127: val terms' = Termtab.fold add_fun terms Symtab.empty boehmes@41127: boehmes@41127: val assms = map (pair ~1) thms @ ithms boehmes@41127: in boehmes@41127: {context=ctxt, typs=typs', terms=terms', rewrite_rules=rules, assms=assms} boehmes@41127: end boehmes@41127: boehmes@41127: boehmes@41127: boehmes@41127: (* preprocessing *) boehmes@41127: boehmes@41127: (** mark built-in constants as Var **) boehmes@41127: boehmes@41127: fun mark_builtins ctxt = boehmes@41127: let boehmes@41127: (* boehmes@41127: Note: schematic terms cannot occur anymore in terms at this stage. boehmes@41127: *) boehmes@41127: fun mark t = boehmes@41127: (case Term.strip_comb t of boehmes@41127: (u as Const (@{const_name If}, _), ts) => marks u ts boehmes@41127: | (u as Const c, ts) => boehmes@41127: (case B.builtin_num ctxt t of boehmes@41127: SOME (n, T) => boehmes@41127: let val v = ((n, 0), T) boehmes@41127: in Vartab.update v #> pair (Var v) end boehmes@41127: | NONE => boehmes@41127: (case B.builtin_fun ctxt c ts of boehmes@41127: SOME ((ni, T), us, U) => boehmes@41127: Vartab.update (ni, U) #> marks (Var (ni, T)) us boehmes@41127: | NONE => marks u ts)) boehmes@41127: | (Abs (n, T, u), ts) => mark u #-> (fn u' => marks (Abs (n, T, u')) ts) boehmes@41127: | (u, ts) => marks u ts) boehmes@41127: boehmes@41127: and marks t ts = fold_map mark ts #>> Term.list_comb o pair t boehmes@41127: boehmes@41127: in (fn ts => swap (fold_map mark ts Vartab.empty)) end boehmes@41127: boehmes@41127: fun mark_builtins' ctxt t = hd (snd (mark_builtins ctxt [t])) boehmes@41127: boehmes@41127: boehmes@41127: (** FIXME **) boehmes@41127: boehmes@41127: local boehmes@41127: (* boehmes@41127: mark constructors and selectors as Vars (forcing eta-expansion), boehmes@41127: add missing datatype selectors via hypothetical definitions, boehmes@41127: also return necessary datatype and record theorems boehmes@41127: *) boehmes@41127: in boehmes@41127: boehmes@41127: fun collect_datatypes_and_records (tr_context, ctxt) ts = boehmes@41127: (([], tr_context, ctxt), ts) boehmes@41127: boehmes@41127: end boehmes@41127: boehmes@41127: boehmes@41127: (** eta-expand quantifiers, let expressions and built-ins *) boehmes@41127: boehmes@41127: local boehmes@41127: fun eta T t = Abs (Name.uu, T, Term.incr_boundvars 1 t $ Bound 0) boehmes@41127: boehmes@41127: fun exp T = eta (Term.domain_type (Term.domain_type T)) boehmes@41127: boehmes@41127: fun exp2 T q = boehmes@41127: let val U = Term.domain_type T boehmes@41127: in Abs (Name.uu, U, q $ eta (Term.domain_type U) (Bound 0)) end boehmes@41127: boehmes@41127: fun exp2' T l = boehmes@41127: let val (U1, U2) = Term.dest_funT T ||> Term.domain_type boehmes@41127: in Abs (Name.uu, U1, eta U2 (l $ Bound 0)) end boehmes@41127: boehmes@41127: fun expf t i T ts = boehmes@41127: let val Ts = U.dest_funT i T |> fst |> drop (length ts) boehmes@41127: in Term.list_comb (t, ts) |> fold_rev eta Ts end boehmes@41127: boehmes@41127: fun expand ((q as Const (@{const_name All}, _)) $ Abs a) = q $ abs_expand a boehmes@41127: | expand ((q as Const (@{const_name All}, T)) $ t) = q $ exp T t boehmes@41127: | expand (q as Const (@{const_name All}, T)) = exp2 T q boehmes@41127: | expand ((q as Const (@{const_name Ex}, _)) $ Abs a) = q $ abs_expand a boehmes@41127: | expand ((q as Const (@{const_name Ex}, T)) $ t) = q $ exp T t boehmes@41127: | expand (q as Const (@{const_name Ex}, T)) = exp2 T q boehmes@41127: | expand ((l as Const (@{const_name Let}, _)) $ t $ Abs a) = boehmes@41127: l $ expand t $ abs_expand a boehmes@41127: | expand ((l as Const (@{const_name Let}, T)) $ t $ u) = boehmes@41127: l $ expand t $ exp (Term.range_type T) u boehmes@41127: | expand ((l as Const (@{const_name Let}, T)) $ t) = exp2 T (l $ expand t) boehmes@41127: | expand (l as Const (@{const_name Let}, T)) = exp2' T l boehmes@41127: | expand (Abs a) = abs_expand a boehmes@41127: | expand t = boehmes@41127: (case Term.strip_comb t of boehmes@41127: (u as Const (@{const_name If}, T), ts) => expf u 3 T (map expand ts) boehmes@41127: | (u as Var ((_, i), T), ts) => expf u i T (map expand ts) boehmes@41127: | (u, ts) => Term.list_comb (u, map expand ts)) boehmes@41127: boehmes@41127: and abs_expand (n, T, t) = Abs (n, T, expand t) boehmes@41127: in boehmes@41127: boehmes@41127: val eta_expand = map expand boehmes@41127: boehmes@41127: end boehmes@41127: boehmes@41127: boehmes@41127: (** lambda-lifting **) boehmes@41127: boehmes@41127: local boehmes@41127: fun mk_def Ts T lhs rhs = boehmes@41127: let boehmes@41127: val eq = HOLogic.eq_const T $ lhs $ rhs boehmes@41127: val trigger = boehmes@41127: [[Const (@{const_name SMT.pat}, T --> @{typ SMT.pattern}) $ lhs]] boehmes@41127: |> map (HOLogic.mk_list @{typ SMT.pattern}) boehmes@41127: |> HOLogic.mk_list @{typ "SMT.pattern list"} boehmes@41127: fun mk_all T t = HOLogic.all_const T $ Abs (Name.uu, T, t) boehmes@41127: in fold mk_all Ts (@{const SMT.trigger} $ trigger $ eq) end boehmes@41127: boehmes@41127: fun replace_lambda Us Ts t (cx as (defs, ctxt)) = boehmes@41127: let boehmes@41127: val T = Term.fastype_of1 (Us @ Ts, t) boehmes@41127: val lev = length Us boehmes@41127: val bs = sort int_ord (Term.add_loose_bnos (t, lev, [])) boehmes@41127: val bss = map_index (fn (i, j) => (j, Integer.add lev i)) bs boehmes@41127: val norm = perhaps (AList.lookup (op =) bss) boehmes@41127: val t' = Term.map_aterms (fn Bound i => Bound (norm i) | t => t) t boehmes@41127: val Ts' = map (nth Ts) bs boehmes@41127: boehmes@41127: fun mk_abs U u = Abs (Name.uu, U, u) boehmes@41127: val abs_rhs = fold mk_abs Ts' (fold mk_abs Us t') boehmes@41127: in boehmes@41127: (case Termtab.lookup defs abs_rhs of boehmes@41127: SOME (f, _) => (Term.list_comb (f, map Bound bs), cx) boehmes@41127: | NONE => boehmes@41127: let boehmes@41127: val (n, ctxt') = boehmes@41127: yield_singleton Variable.variant_fixes Name.uu ctxt boehmes@41127: val f = Free (n, rev Ts' ---> (rev Us ---> T)) boehmes@41127: fun mk_bapp i t = t $ Bound i boehmes@41127: val lhs = boehmes@41127: f boehmes@41127: |> fold_rev (mk_bapp o snd) bss boehmes@41127: |> fold_rev mk_bapp (0 upto (length Us - 1)) boehmes@41127: val def = mk_def (Us @ Ts') T lhs t' boehmes@41127: in (f, (Termtab.update (abs_rhs, (f, def)) defs, ctxt')) end) boehmes@41127: end boehmes@41127: boehmes@41127: fun dest_abs Ts (Abs (_, T, t)) = dest_abs (T :: Ts) t boehmes@41127: | dest_abs Ts t = (Ts, t) boehmes@41127: boehmes@41127: fun traverse Ts t = boehmes@41127: (case t of boehmes@41127: (q as Const (@{const_name All}, _)) $ Abs a => boehmes@41127: abs_traverse Ts a #>> (fn a' => q $ Abs a') boehmes@41127: | (q as Const (@{const_name Ex}, _)) $ Abs a => boehmes@41127: abs_traverse Ts a #>> (fn a' => q $ Abs a') boehmes@41127: | (l as Const (@{const_name Let}, _)) $ u $ Abs a => boehmes@41127: traverse Ts u ##>> abs_traverse Ts a #>> boehmes@41127: (fn (u', a') => l $ u' $ Abs a') boehmes@41127: | Abs _ => boehmes@41127: let val (Us, u) = dest_abs [] t boehmes@41127: in traverse (Us @ Ts) u #-> replace_lambda Us Ts end boehmes@41127: | u1 $ u2 => traverse Ts u1 ##>> traverse Ts u2 #>> (op $) boehmes@41127: | _ => pair t) boehmes@41127: boehmes@41127: and abs_traverse Ts (n, T, t) = traverse (T::Ts) t #>> (fn t' => (n, T, t')) boehmes@41127: in boehmes@41127: boehmes@41127: fun lift_lambdas ctxt ts = boehmes@41127: (Termtab.empty, ctxt) boehmes@41127: |> fold_map (traverse []) ts boehmes@41127: |> (fn (us, (defs, ctxt')) => boehmes@41127: (ctxt', Termtab.fold (cons o snd o snd) defs us)) boehmes@41127: boehmes@41127: end boehmes@41127: boehmes@41127: boehmes@41127: (** introduce explicit applications **) boehmes@41127: boehmes@41127: local boehmes@41127: (* boehmes@41127: Make application explicit for functions with varying number of arguments. boehmes@41127: *) boehmes@41127: boehmes@41127: fun add t ts = boehmes@41127: Termtab.map_default (t, []) (Ord_List.insert int_ord (length ts)) boehmes@41127: boehmes@41127: fun collect t = boehmes@41127: (case Term.strip_comb t of boehmes@41127: (u as Const _, ts) => add u ts #> fold collect ts boehmes@41127: | (u as Free _, ts) => add u ts #> fold collect ts boehmes@41127: | (Abs (_, _, u), ts) => collect u #> fold collect ts boehmes@41127: | (_, ts) => fold collect ts) boehmes@41127: boehmes@41127: fun app ts (t, T) = boehmes@41127: let val f = Const (@{const_name SMT.fun_app}, T --> T) boehmes@41127: in (Term.list_comb (f $ t, ts), snd (U.dest_funT (length ts) T)) end boehmes@41127: boehmes@41127: fun appl _ _ [] = fst boehmes@41127: | appl _ [] ts = fst o app ts boehmes@41127: | appl i (k :: ks) ts = boehmes@41127: let val (ts1, ts2) = chop (k - i) ts boehmes@41127: in appl k ks ts2 o app ts1 end boehmes@41127: boehmes@41127: fun appl0 [_] ts (t, _) = Term.list_comb (t, ts) boehmes@41127: | appl0 (0 :: ks) ts tT = appl 0 ks ts tT boehmes@41127: | appl0 ks ts tT = appl 0 ks ts tT boehmes@41127: boehmes@41127: fun apply terms T t ts = appl0 (Termtab.lookup_list terms t) ts (t, T) boehmes@41127: boehmes@41127: fun get_arities i t = boehmes@41127: (case Term.strip_comb t of boehmes@41127: (Bound j, ts) => boehmes@41127: (if i = j then Ord_List.insert int_ord (length ts) else I) #> boehmes@41127: fold (get_arities i) ts boehmes@41127: | (Abs (_, _, u), ts) => get_arities (i+1) u #> fold (get_arities i) ts boehmes@41127: | (_, ts) => fold (get_arities i) ts) boehmes@41127: in boehmes@41127: boehmes@41127: fun intro_explicit_application ts = boehmes@41127: let boehmes@41127: val terms = fold collect ts Termtab.empty boehmes@41127: boehmes@41127: fun traverse (env as (arities, Ts)) t = boehmes@41127: (case Term.strip_comb t of boehmes@41127: (u as Const (_, T), ts) => apply terms T u (map (traverse env) ts) boehmes@41127: | (u as Free (_, T), ts) => apply terms T u (map (traverse env) ts) boehmes@41127: | (u as Bound i, ts) => boehmes@41127: appl0 (nth arities i) (map (traverse env) ts) (u, nth Ts i) boehmes@41127: | (Abs (n, T, u), ts) => boehmes@41127: let val env' = (get_arities 0 u [] :: arities, T :: Ts) boehmes@41127: in traverses env (Abs (n, T, traverse env' u)) ts end boehmes@41127: | (u, ts) => traverses env u ts) boehmes@41127: and traverses env t ts = Term.list_comb (t, map (traverse env) ts) boehmes@41127: in map (traverse ([], [])) ts end boehmes@41127: boehmes@41127: val fun_app_eq = mk_meta_eq @{thm SMT.fun_app_def} boehmes@41127: boehmes@41127: end boehmes@41127: boehmes@41127: boehmes@41127: (** map HOL formulas to FOL formulas (i.e., separate formulas froms terms) **) boehmes@41127: boehmes@41127: val tboolT = @{typ SMT.term_bool} boehmes@41127: val term_true = Const (@{const_name True}, tboolT) boehmes@41127: val term_false = Const (@{const_name False}, tboolT) boehmes@41127: boehmes@41127: val term_bool = @{lemma "True ~= False" by simp} boehmes@41127: val term_bool_prop = boehmes@41127: let boehmes@41127: fun replace @{const HOL.eq (bool)} = @{const HOL.eq (SMT.term_bool)} boehmes@41127: | replace @{const True} = term_true boehmes@41127: | replace @{const False} = term_false boehmes@41127: | replace t = t boehmes@41127: in boehmes@41127: Term.map_aterms replace (HOLogic.dest_Trueprop (Thm.prop_of term_bool)) boehmes@41127: end boehmes@41127: boehmes@41127: val fol_rules = [ boehmes@41127: Let_def, boehmes@41127: @{lemma "P = True == P" by (rule eq_reflection) simp}, boehmes@41127: @{lemma "if P then True else False == P" by (rule eq_reflection) simp}] boehmes@41127: boehmes@41127: fun reduce_let (Const (@{const_name Let}, _) $ t $ u) = boehmes@41127: reduce_let (Term.betapply (u, t)) boehmes@41127: | reduce_let (t $ u) = reduce_let t $ reduce_let u boehmes@41127: | reduce_let (Abs (n, T, t)) = Abs (n, T, reduce_let t) boehmes@41127: | reduce_let t = t boehmes@41127: boehmes@41127: fun is_pred_type NONE = false boehmes@41127: | is_pred_type (SOME T) = (Term.body_type T = @{typ bool}) boehmes@41127: boehmes@41127: fun is_conn_type NONE = false boehmes@41127: | is_conn_type (SOME T) = boehmes@41127: forall (equal @{typ bool}) (Term.body_type T :: Term.binder_types T) boehmes@41127: boehmes@41127: fun revert_typ @{typ SMT.term_bool} = @{typ bool} boehmes@41127: | revert_typ (Type (n, Ts)) = Type (n, map revert_typ Ts) boehmes@41127: | revert_typ T = T boehmes@41127: boehmes@41127: val revert_types = Term.map_types revert_typ boehmes@41127: boehmes@41127: fun folify ctxt builtins = boehmes@41127: let boehmes@41127: fun as_term t = @{const HOL.eq (SMT.term_bool)} $ t $ term_true boehmes@41127: boehmes@41127: fun as_tbool @{typ bool} = tboolT boehmes@41127: | as_tbool (Type (n, Ts)) = Type (n, map as_tbool Ts) boehmes@41127: | as_tbool T = T boehmes@41127: fun mapTs f g i = U.dest_funT i #> (fn (Ts, T) => map f Ts ---> g T) boehmes@41127: fun predT i = mapTs as_tbool I i boehmes@41127: fun funcT i = mapTs as_tbool as_tbool i boehmes@41127: fun func i (n, T) = (n, funcT i T) boehmes@41127: boehmes@41127: fun map_ifT T = T |> Term.dest_funT ||> funcT 2 |> (op -->) boehmes@41127: val if_term = @{const If (bool)} |> Term.dest_Const ||> map_ifT |> Const boehmes@41127: fun wrap_in_if t = if_term $ t $ term_true $ term_false boehmes@41127: boehmes@41127: fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t)) boehmes@41127: boehmes@41127: fun in_term t = boehmes@41127: (case Term.strip_comb t of boehmes@41127: (Const (n as @{const_name If}, T), [t1, t2, t3]) => boehmes@41127: Const (n, map_ifT T) $ in_form t1 $ in_term t2 $ in_term t3 boehmes@41127: | (Const (@{const_name HOL.eq}, _), _) => wrap_in_if (in_form t) boehmes@41127: | (Var (ni as (_, i), T), ts) => boehmes@41127: let val U = Vartab.lookup builtins ni boehmes@41127: in boehmes@41127: if is_conn_type U orelse is_pred_type U then wrap_in_if (in_form t) boehmes@41127: else Term.list_comb (Var (ni, funcT i T), map in_term ts) boehmes@41127: end boehmes@41127: | (Const c, ts) => boehmes@41127: Term.list_comb (Const (func (length ts) c), map in_term ts) boehmes@41127: | (Free c, ts) => boehmes@41127: Term.list_comb (Free (func (length ts) c), map in_term ts) boehmes@41127: | _ => t) boehmes@41127: boehmes@41127: and in_weight ((c as @{const SMT.weight}) $ w $ t) = c $ w $ in_form t boehmes@41127: | in_weight t = in_form t boehmes@41127: boehmes@41127: and in_pat (Const (c as (@{const_name pat}, _)) $ t) = boehmes@41127: Const (func 1 c) $ in_term t boehmes@41127: | in_pat (Const (c as (@{const_name nopat}, _)) $ t) = boehmes@41127: Const (func 1 c) $ in_term t boehmes@41127: | in_pat t = raise TERM ("bad pattern", [t]) boehmes@41127: boehmes@41127: and in_pats ps = boehmes@41127: in_list @{typ "pattern list"} (in_list @{typ pattern} in_pat) ps boehmes@41127: boehmes@41127: and in_trig ((c as @{const trigger}) $ p $ t) = c $ in_pats p $ in_weight t boehmes@41127: | in_trig t = in_weight t boehmes@41127: boehmes@41127: and in_form t = boehmes@41127: (case Term.strip_comb t of boehmes@41127: (q as Const (qn, _), [Abs (n, T, u)]) => boehmes@41127: if member (op =) [@{const_name All}, @{const_name Ex}] qn then boehmes@41127: q $ Abs (n, as_tbool T, in_trig u) boehmes@41127: else as_term (in_term t) boehmes@41127: | (u as Const (@{const_name If}, _), ts) => boehmes@41127: Term.list_comb (u, map in_form ts) boehmes@41127: | (b as @{const HOL.eq (bool)}, ts) => Term.list_comb (b, map in_form ts) boehmes@41127: | (Const (n as @{const_name HOL.eq}, T), ts) => boehmes@41127: Term.list_comb (Const (n, predT 2 T), map in_term ts) boehmes@41127: | (b as Var (ni as (_, i), T), ts) => boehmes@41127: if is_conn_type (Vartab.lookup builtins ni) then boehmes@41127: Term.list_comb (b, map in_form ts) boehmes@41127: else if is_pred_type (Vartab.lookup builtins ni) then boehmes@41127: Term.list_comb (Var (ni, predT i T), map in_term ts) boehmes@41127: else as_term (in_term t) boehmes@41127: | _ => as_term (in_term t)) boehmes@41127: in boehmes@41127: map (reduce_let #> in_form) #> boehmes@41127: cons (mark_builtins' ctxt term_bool_prop) #> boehmes@41127: pair (fol_rules, [term_bool]) boehmes@41127: end boehmes@41127: boehmes@41127: boehmes@41127: boehmes@41127: (* translation into intermediate format *) boehmes@41127: boehmes@41127: (** utility functions **) boehmes@36898: boehmes@36898: val quantifier = (fn boehmes@36898: @{const_name All} => SOME SForall boehmes@36898: | @{const_name Ex} => SOME SExists boehmes@36898: | _ => NONE) boehmes@36898: boehmes@36898: fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) = boehmes@36898: if q = qname then group_quant qname (T :: Ts) u else (Ts, t) boehmes@36898: | group_quant _ Ts t = (Ts, t) boehmes@36898: boehmes@40664: fun dest_weight (@{const SMT.weight} $ w $ t) = boehmes@40664: (SOME (snd (HOLogic.dest_number w)), t) boehmes@40664: | dest_weight t = (NONE, t) boehmes@40664: boehmes@37124: fun dest_pat (Const (@{const_name pat}, _) $ t) = (t, true) boehmes@37124: | dest_pat (Const (@{const_name nopat}, _) $ t) = (t, false) boehmes@41127: | dest_pat t = raise TERM ("bad pattern", [t]) boehmes@37124: boehmes@37124: fun dest_pats [] = I boehmes@37124: | dest_pats ts = boehmes@37124: (case map dest_pat ts |> split_list ||> distinct (op =) of boehmes@37124: (ps, [true]) => cons (SPat ps) boehmes@37124: | (ps, [false]) => cons (SNoPat ps) boehmes@41127: | _ => raise TERM ("bad multi-pattern", ts)) boehmes@36898: boehmes@40579: fun dest_trigger (@{const trigger} $ tl $ t) = boehmes@37124: (rev (fold (dest_pats o HOLogic.dest_list) (HOLogic.dest_list tl) []), t) boehmes@36898: | dest_trigger t = ([], t) boehmes@36898: boehmes@36898: fun dest_quant qn T t = quantifier qn |> Option.map (fn q => boehmes@36898: let boehmes@36898: val (Ts, u) = group_quant qn [T] t boehmes@40664: val (ps, p) = dest_trigger u boehmes@40664: val (w, b) = dest_weight p boehmes@40664: in (q, rev Ts, ps, w, b) end) boehmes@36898: boehmes@36898: fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat boehmes@36898: | fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat boehmes@36898: boehmes@36898: boehmes@41127: (** translation from Isabelle terms into SMT intermediate terms **) boehmes@36898: boehmes@41127: fun intermediate header dtyps ctxt ts trx = boehmes@41059: let boehmes@41127: fun transT (T as TFree _) = add_typ T true boehmes@41127: | transT (T as TVar _) = (fn _ => raise TYPE ("bad SMT type", [T], [])) boehmes@41127: | transT (T as Type _) = boehmes@41059: (case B.builtin_typ ctxt T of boehmes@39298: SOME n => pair n boehmes@41127: | NONE => add_typ T true) boehmes@36898: boehmes@41127: val unmarked_builtins = [@{const_name If}, @{const_name HOL.eq}] boehmes@36898: boehmes@36898: fun app n ts = SApp (n, ts) boehmes@36898: boehmes@36898: fun trans t = boehmes@36898: (case Term.strip_comb t of boehmes@36898: (Const (qn, _), [Abs (_, T, t1)]) => boehmes@36898: (case dest_quant qn T t1 of boehmes@40664: SOME (q, Ts, ps, w, b) => boehmes@36898: fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>> boehmes@40664: trans b #>> (fn ((Ts', ps'), b') => SQua (q, Ts', ps', w, b')) boehmes@41127: | NONE => raise TERM ("unsupported quantifier", [t])) boehmes@36898: | (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) => boehmes@36898: transT T ##>> trans t1 ##>> trans t2 #>> boehmes@36898: (fn ((U, u1), u2) => SLet (U, u1, u2)) boehmes@41127: | (Var ((n, _), _), ts) => fold_map trans ts #>> app n boehmes@41127: | (u as Const (c as (n, T)), ts) => boehmes@41127: if member (op =) unmarked_builtins n then boehmes@41127: (case B.builtin_fun ctxt c ts of boehmes@41127: SOME (((m, _), _), us, _) => fold_map trans us #>> app m boehmes@41127: | NONE => raise TERM ("not a built-in symbol", [t])) boehmes@41127: else transs u T ts boehmes@41127: | (u as Free (_, T), ts) => transs u T ts boehmes@36898: | (Bound i, []) => pair (SVar i) boehmes@41127: | _ => raise TERM ("bad SMT term", [t])) boehmes@41127: boehmes@36898: and transs t T ts = boehmes@40663: let val (Us, U) = U.dest_funT (length ts) T boehmes@36898: in boehmes@36898: fold_map transT Us ##>> transT U #-> (fn Up => boehmes@41127: add_fun t (SOME Up) ##>> fold_map trans ts #>> SApp) boehmes@36898: end boehmes@41127: boehmes@41127: val (us, trx') = fold_map trans ts trx boehmes@41127: in ((sign_of (header ts) dtyps trx', us), trx') end boehmes@41127: boehmes@41127: boehmes@41127: boehmes@41127: (* translation *) boehmes@41127: boehmes@41127: structure Configs = Generic_Data boehmes@41127: ( boehmes@41127: type T = (Proof.context -> config) U.dict boehmes@41127: val empty = [] boehmes@41127: val extend = I boehmes@41127: val merge = U.dict_merge fst boehmes@41127: ) boehmes@41127: boehmes@41127: fun add_config (cs, cfg) = Configs.map (U.dict_update (cs, cfg)) boehmes@41127: boehmes@41127: fun translate ctxt comments ithms = boehmes@41127: let boehmes@41127: val cs = SMT_Config.solver_class_of ctxt boehmes@41127: val {prefixes, is_fol, header, has_datatypes, serialize} = boehmes@41127: (case U.dict_get (Configs.get (Context.Proof ctxt)) cs of boehmes@41127: SOME cfg => cfg ctxt boehmes@41127: | NONE => error ("SMT: no translation configuration found " ^ boehmes@41127: "for solver class " ^ quote (U.string_of_class cs))) boehmes@41127: boehmes@41127: val with_datatypes = boehmes@41127: has_datatypes andalso Config.get ctxt SMT_Config.datatypes boehmes@41127: boehmes@41127: fun no_dtyps (tr_context, ctxt) ts = (([], tr_context, ctxt), ts) boehmes@41127: boehmes@41127: val (builtins, ts1) = boehmes@41127: ithms boehmes@41127: |> map (HOLogic.dest_Trueprop o Thm.prop_of o snd) boehmes@41127: |> map (Envir.eta_contract o Envir.beta_norm) boehmes@41127: |> mark_builtins ctxt boehmes@41127: boehmes@41127: val ((dtyps, tr_context, ctxt1), ts2) = boehmes@41127: ((make_tr_context prefixes, ctxt), ts1) boehmes@41127: |-> (if with_datatypes then collect_datatypes_and_records else no_dtyps) boehmes@41127: boehmes@41127: val (ctxt2, ts3) = boehmes@41127: ts2 boehmes@41127: |> eta_expand boehmes@41127: |> lift_lambdas ctxt1 boehmes@41127: ||> intro_explicit_application boehmes@41127: boehmes@41127: val ((rewrite_rules, extra_thms), ts4) = boehmes@41127: (if is_fol then folify ctxt2 builtins else pair ([], [])) ts3 boehmes@41127: boehmes@41127: val rewrite_rules' = fun_app_eq :: rewrite_rules boehmes@36898: in boehmes@41127: (ts4, tr_context) boehmes@41127: |-> intermediate header dtyps ctxt2 boehmes@41127: |>> uncurry (serialize comments) boehmes@41127: ||> recon_of ctxt2 rewrite_rules' extra_thms ithms revert_typ revert_types boehmes@36898: end boehmes@36898: boehmes@36898: end