src/HOL/Tools/SMT/smt_translate.ML
author boehmes
Sun, 19 Dec 2010 18:54:29 +0100
changeset 41281 679118e35378
parent 41250 41f86829e22f
child 41328 6792a5c92a58
permissions -rw-r--r--
removed odd decoration of built-in symbols as Vars (instead provide built-in desctructor functions along with their inverse functions); removed odd retyping during folify (instead, keep all terms well-typed)

(*  Title:      HOL/Tools/SMT/smt_translate.ML
    Author:     Sascha Boehme, TU Muenchen

Translate theorems into an SMT intermediate format and serialize them.
*)

signature SMT_TRANSLATE =
sig
  (*intermediate term structure*)
  datatype squant = SForall | SExists
  datatype 'a spattern = SPat of 'a list | SNoPat of 'a list
  datatype sterm =
    SVar of int |
    SApp of string * sterm list |
    SLet of string * sterm * sterm |
    SQua of squant * string list * sterm spattern list * int option * sterm

  (*translation configuration*)
  type prefixes = {sort_prefix: string, func_prefix: string}
  type sign = {
    header: string list,
    sorts: string list,
    dtyps: (string * (string * (string * string) list) list) list list,
    funcs: (string * (string list * string)) list }
  type config = {
    prefixes: prefixes,
    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,
    rewrite_rules: thm list,
    assms: (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

structure SMT_Translate: SMT_TRANSLATE =
struct

structure U = SMT_Utils
structure B = SMT_Builtin


(* intermediate term structure *)

datatype squant = SForall | SExists

datatype 'a spattern = SPat of 'a list | SNoPat of 'a list

datatype sterm =
  SVar of int |
  SApp of string * sterm list |
  SLet of string * sterm * sterm |
  SQua of squant * string list * sterm spattern list * int option * sterm



(* translation configuration *)

type prefixes = {sort_prefix: string, func_prefix: string}

type sign = {
  header: string list,
  sorts: string list,
  dtyps: (string * (string * (string * string) list) list) list list,
  funcs: (string * (string list * string)) list }

type config = {
  prefixes: prefixes,
  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,
  rewrite_rules: thm list,
  assms: (int * thm) list }



(* 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 (_, _, typs, _, _, terms) =
  let
    fun add_typ (T, (n, _)) = Symtab.update (n, T)
    val typs' = Typtab.fold add_typ typs Symtab.empty

    fun add_fun (t, (n, _)) = Symtab.update (n, 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 *)

(** FIXME **)

local
  (*
    force eta-expansion for constructors and selectors,
    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 k i T t =
    let val Ts = drop i (fst (U.dest_funT k T))
    in
      Term.incr_boundvars (length Ts) t
      |> fold_index (fn (i, _) => fn u => u $ Bound i) Ts
      |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts
    end
in

fun eta_expand ctxt =
  let
    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 t =
          (case Term.strip_comb t of
            (u as Const (c as (_, T)), ts) =>
              (case B.dest_builtin ctxt c ts of
                SOME (_, k, us, mk) =>
                  if k = length us then mk (map expand us)
                  else expf k (length ts) T (mk (map expand us))
              | NONE => Term.list_comb (u, map expand ts))
          | (Abs a, ts) => Term.list_comb (abs_expand a, map expand ts)
          | (u, ts) => Term.list_comb (u, map expand ts))

    and abs_expand (n, T, t) = Abs (n, T, expand t)
  
  in map expand end

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 mk_abs Ts = fold (fn T => fn t => Abs (Name.uu, T, t)) Ts

  fun dest_abs Ts (Abs (_, T, t)) = dest_abs (T :: Ts) t
    | dest_abs Ts t = (Ts, t)

  fun replace_lambda Us Ts t (cx as (defs, ctxt)) =
    let
      val t1 = mk_abs Us t
      val bs = sort int_ord (Term.add_loose_bnos (t1, 0, []))
      fun rep i k = if member (op =) bs i then (Bound k, k+1) else (Bound i, k)
      val (rs, _) = fold_map rep (0 upto length Ts - 1) 0
      val t2 = Term.subst_bounds (rs, t1)
      val Ts' = map (nth Ts) bs 
      val (_, t3) = dest_abs [] t2
      val t4 = mk_abs Ts' t2

      val T = Term.fastype_of1 (Us @ Ts, t)
      fun app f = Term.list_comb (f, map Bound (rev bs))
    in
      (case Termtab.lookup defs t4 of
        SOME (f, _) => (app f, cx)
      | NONE =>
          let
            val (n, ctxt') =
              yield_singleton Variable.variant_fixes Name.uu ctxt
            val (is, UTs) = split_list (map_index I (Us @ Ts'))
            val f = Free (n, rev UTs ---> T)
            val lhs = Term.list_comb (f, map Bound (rev is))
            val def = mk_def UTs (Term.fastype_of1 (Us @ Ts, t)) lhs t3
          in (app f, (Termtab.update (t4, (f, def)) defs, ctxt')) end)
    end

  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 i = Termtab.map_default (t, i) (Integer.min i)

  fun min_arities t =
    (case Term.strip_comb t of
      (u as Const _, ts) => add u (length ts) #> fold min_arities ts
    | (u as Free _, ts) => add u (length ts) #> fold min_arities ts
    | (Abs (_, _, u), ts) => min_arities u #> fold min_arities ts
    | (_, ts) => fold min_arities ts)

  fun app u (t, T) =
    (Const (@{const_name SMT.fun_app}, T --> T) $ t $ u, Term.range_type T)

  fun apply i t T ts =
    let val (ts1, ts2) = chop i ts
    in fst (fold app ts2 (Term.list_comb (t, ts1), snd (U.dest_funT i T))) end
in

fun intro_explicit_application ts =
  let
    val arities = fold min_arities ts Termtab.empty
    fun apply' t = apply (the (Termtab.lookup arities t)) t

    fun traverse Ts t =
      (case Term.strip_comb t of
        (u as Const (_, T), ts) => apply' u T (map (traverse Ts) ts)
      | (u as Free (_, T), ts) => apply' u T (map (traverse Ts) ts)
      | (u as Bound i, ts) => apply 0 u (nth Ts i) (map (traverse Ts) ts)
      | (Abs (n, T, u), ts) => traverses Ts (Abs (n, T, traverse (T::Ts) u)) ts
      | (u, ts) => traverses Ts u ts)
    and traverses Ts t ts = Term.list_comb (t, map (traverse Ts) 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) **)

local
  val term_bool = @{lemma "SMT.term_true ~= SMT.term_false"
    by (simp add: SMT.term_true_def SMT.term_false_def)}

  val fol_rules = [
    Let_def,
    mk_meta_eq @{thm SMT.term_true_def},
    mk_meta_eq @{thm SMT.term_false_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 as_term t = @{const HOL.eq (bool)} $ t $ @{const SMT.term_true}

  fun wrap_in_if t =
    @{const If (bool)} $ t $ @{const SMT.term_true} $ @{const SMT.term_false}

  fun is_builtin_conn_or_pred ctxt c ts =
    is_some (B.dest_builtin_conn ctxt c ts) orelse
    is_some (B.dest_builtin_pred ctxt c ts)

  fun builtin b ctxt c ts =
    (case (Const c, ts) of
      (@{const HOL.eq (bool)}, [t, u]) =>
        if t = @{const SMT.term_true} orelse u = @{const SMT.term_true} then
          B.dest_builtin_eq ctxt t u
        else b ctxt c ts
    | _ => b ctxt c ts)
in

fun folify ctxt =
  let
    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 True}, []) => @{const SMT.term_true}
      | (@{const False}, []) => @{const SMT.term_false}
      | (u as Const (@{const_name If}, _), [t1, t2, t3]) =>
          u $ in_form t1 $ in_term t2 $ in_term t3
      | (Const c, ts) =>
          if is_builtin_conn_or_pred ctxt c ts then wrap_in_if (in_form t)
          else Term.list_comb (Const c, map in_term ts)
      | (Free c, ts) => Term.list_comb (Free 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 ((p as Const (@{const_name SMT.pat}, _)) $ t) = p $ in_term t
      | in_pat ((p as Const (@{const_name SMT.nopat}, _)) $ t) = p $ in_term t
      | in_pat t = raise TERM ("bad pattern", [t])

    and in_pats ps =
      in_list @{typ "SMT.pattern list"} (in_list @{typ SMT.pattern} in_pat) ps

    and in_trigger ((c as @{const SMT.trigger}) $ p $ t) =
          c $ in_pats p $ in_weight t
      | in_trigger 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, T, in_trigger u)
          else as_term (in_term t)
      | (Const c, ts) =>
          (case B.dest_builtin_conn ctxt c ts of
            SOME (_, _, us, mk) => mk (map in_form us)
          | NONE =>
              (case B.dest_builtin_pred ctxt c ts of
                SOME (_, _, us, mk) => mk (map in_term us)
              | NONE => as_term (in_term t)))
      | _ => as_term (in_term t))
  in
    map (reduce_let #> in_form) #>
    cons (U.prop_of term_bool) #>
    pair (fol_rules, [term_bool], builtin)
  end

end


(* translation into intermediate format *)

(** utility functions **)

val quantifier = (fn
    @{const_name All} => SOME SForall
  | @{const_name Ex} => SOME SExists
  | _ => NONE)

fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) =
      if q = qname then group_quant qname (T :: Ts) u else (Ts, t)
  | group_quant _ Ts t = (Ts, t)

fun dest_weight (@{const SMT.weight} $ w $ t) =
      (SOME (snd (HOLogic.dest_number w)), t)
  | dest_weight t = (NONE, t)

fun dest_pat (Const (@{const_name SMT.pat}, _) $ t) = (t, true)
  | dest_pat (Const (@{const_name SMT.nopat}, _) $ t) = (t, false)
  | 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 ("bad multi-pattern", ts))

fun dest_trigger (@{const SMT.trigger} $ tl $ t) =
      (rev (fold (dest_pats o HOLogic.dest_list) (HOLogic.dest_list tl) []), t)
  | dest_trigger t = ([], t)

fun dest_quant qn T t = quantifier qn |> Option.map (fn q =>
  let
    val (Ts, u) = group_quant qn [T] t
    val (ps, p) = dest_trigger u
    val (w, b) = dest_weight p
  in (q, rev Ts, ps, w, b) end)

fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat
  | fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat


(** translation from Isabelle terms into SMT intermediate terms **)

fun intermediate header dtyps builtin ctxt ts trx =
  let
    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.dest_builtin_typ ctxt T of
            SOME n => pair n
          | NONE => add_typ T true)

    fun app n ts = SApp (n, ts)

    fun trans t =
      (case Term.strip_comb t of
        (Const (qn, _), [Abs (_, T, t1)]) =>
          (case dest_quant qn T t1 of
            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 ("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))
      | (u as Const (c as (_, T)), ts) =>
          (case builtin ctxt c ts of
            SOME (n, _, us, _) => fold_map trans us #>> app n
          | NONE => transs u T ts)
      | (u as Free (_, T), ts) => transs u T ts
      | (Bound i, []) => pair (SVar i)
      | _ => 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 =>
        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
  fun merge data = U.dict_merge fst data
)

fun add_config (cs, cfg) = Configs.map (U.dict_update (cs, cfg))

fun get_config ctxt = 
  let val cs = SMT_Config.solver_class_of ctxt
  in
    (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)))
  end

fun translate ctxt comments ithms =
  let
    val {prefixes, is_fol, header, has_datatypes, serialize} = get_config ctxt

    val with_datatypes =
      has_datatypes andalso Config.get ctxt SMT_Config.datatypes

    fun no_dtyps (tr_context, ctxt) ts = (([], tr_context, ctxt), ts)

    val ts1 = map (Envir.beta_eta_contract o U.prop_of o snd) ithms

    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 ctxt1
      |> lift_lambdas ctxt1
      ||> intro_explicit_application

    val ((rewrite_rules, extra_thms, builtin), ts4) =
      (if is_fol then folify ctxt2 else pair ([], [], I)) ts3

    val rewrite_rules' = fun_app_eq :: rewrite_rules
  in
    (ts4, tr_context)
    |-> intermediate header dtyps (builtin B.dest_builtin) ctxt2
    |>> uncurry (serialize comments)
    ||> recon_of ctxt2 rewrite_rules' extra_thms ithms
  end

end