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@36898: (* 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@36898: (* configuration options *) 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@41059: header: Proof.context -> 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@36898: typs: typ Symtab.table, boehmes@36898: terms: term Symtab.table, boehmes@36898: unfolds: thm list, boehmes@40161: assms: (int * thm) list } boehmes@36898: boehmes@40161: val translate: config -> 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@36898: (* configuration options *) 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@41059: header: Proof.context -> 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@36898: typs: typ Symtab.table, boehmes@36898: terms: term Symtab.table, boehmes@36898: unfolds: thm list, boehmes@40161: assms: (int * thm) list } boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36898: (* 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@37124: | dest_pat t = raise TERM ("dest_pat", [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@37124: | _ => raise TERM ("dest_pats", 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: fun prop_of thm = HOLogic.dest_Trueprop (Thm.prop_of thm) boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41059: (* map HOL formulas to FOL formulas (i.e., separate formulas froms terms) *) boehmes@36898: boehmes@41059: val tboolT = @{typ SMT.term_bool} boehmes@41059: val term_true = Const (@{const_name True}, tboolT) boehmes@41059: val term_false = Const (@{const_name False}, tboolT) boehmes@36898: boehmes@41059: val term_bool = @{lemma "True ~= False" by simp} boehmes@41059: val term_bool_prop = boehmes@41059: let boehmes@41059: fun replace @{const HOL.eq (bool)} = @{const HOL.eq (SMT.term_bool)} boehmes@41059: | replace @{const True} = term_true boehmes@41059: | replace @{const False} = term_false boehmes@41059: | replace t = t boehmes@41059: in Term.map_aterms replace (prop_of term_bool) end boehmes@36898: boehmes@36898: val needs_rewrite = Thm.prop_of #> Term.exists_subterm (fn boehmes@36898: Const (@{const_name Let}, _) => true boehmes@40579: | @{const HOL.eq (bool)} $ _ $ @{const True} => true boehmes@40579: | Const (@{const_name If}, _) $ _ $ @{const True} $ @{const False} => true boehmes@36898: | _ => false) boehmes@36898: boehmes@36898: val rewrite_rules = [ boehmes@36898: Let_def, boehmes@36898: @{lemma "P = True == P" by (rule eq_reflection) simp}, boehmes@36898: @{lemma "if P then True else False == P" by (rule eq_reflection) simp}] boehmes@36898: boehmes@41059: fun rewrite ctxt ct = boehmes@41059: Conv.top_sweep_conv (fn ctxt' => boehmes@41059: Conv.rewrs_conv rewrite_rules then_conv rewrite ctxt') ctxt ct boehmes@36898: boehmes@36898: fun normalize ctxt thm = boehmes@36898: if needs_rewrite thm then Conv.fconv_rule (rewrite ctxt) thm else thm boehmes@36898: boehmes@41059: fun revert_typ @{typ SMT.term_bool} = @{typ bool} boehmes@41059: | revert_typ (Type (n, Ts)) = Type (n, map revert_typ Ts) boehmes@41059: | revert_typ T = T boehmes@36898: boehmes@41059: val revert_types = Term.map_types revert_typ boehmes@36898: boehmes@41059: fun folify ctxt = boehmes@36898: let boehmes@41059: fun is_builtin_conn (@{const_name True}, _) _ = false boehmes@41059: | is_builtin_conn (@{const_name False}, _) _ = false boehmes@41059: | is_builtin_conn c ts = B.is_builtin_conn ctxt c ts boehmes@36898: boehmes@41059: fun as_term t = @{const HOL.eq (SMT.term_bool)} $ t $ term_true boehmes@40681: boehmes@41059: fun as_tbool @{typ bool} = tboolT boehmes@41059: | as_tbool (Type (n, Ts)) = Type (n, map as_tbool Ts) boehmes@41059: | as_tbool T = T boehmes@36898: fun mapTs f g = Term.strip_type #> (fn (Ts, T) => map f Ts ---> g T) boehmes@41059: fun predT T = mapTs as_tbool I T boehmes@41059: fun funcT T = mapTs as_tbool as_tbool T boehmes@41059: fun func (n, T) = Const (n, funcT T) boehmes@36898: boehmes@41059: fun map_ifT T = T |> Term.dest_funT ||> funcT |> (op -->) boehmes@41059: val if_term = @{const If (bool)} |> Term.dest_Const ||> map_ifT |> Const boehmes@41059: fun wrap_in_if t = if_term $ t $ term_true $ term_false boehmes@36898: boehmes@36898: fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t)) boehmes@36898: boehmes@36898: fun in_term t = boehmes@36898: (case Term.strip_comb t of boehmes@41059: (Const (c as @{const_name If}, T), [t1, t2, t3]) => boehmes@41059: Const (c, map_ifT T) $ in_form t1 $ in_term t2 $ in_term t3 boehmes@41059: | (Const c, ts) => boehmes@41059: if is_builtin_conn c ts orelse B.is_builtin_pred ctxt c ts boehmes@36898: then wrap_in_if (in_form t) boehmes@41059: else Term.list_comb (func c, map in_term ts) boehmes@41059: | (Free (n, T), ts) => Term.list_comb (Free (n, funcT T), map in_term ts) boehmes@36898: | _ => t) boehmes@36898: boehmes@40664: and in_weight ((c as @{const SMT.weight}) $ w $ t) = c $ w $ in_form t boehmes@40664: | in_weight t = in_form t boehmes@40664: boehmes@41059: and in_pat (Const (c as (@{const_name pat}, _)) $ t) = func c $ in_term t boehmes@41059: | in_pat (Const (c as (@{const_name nopat}, _)) $ t) = func c $ in_term t boehmes@36898: | in_pat t = raise TERM ("in_pat", [t]) boehmes@36898: boehmes@37124: and in_pats ps = boehmes@37124: in_list @{typ "pattern list"} (in_list @{typ pattern} in_pat) ps boehmes@36898: boehmes@40664: and in_trig ((c as @{const trigger}) $ p $ t) = c $ in_pats p $ in_weight t boehmes@40664: | in_trig t = in_weight t boehmes@36898: boehmes@36898: and in_form t = boehmes@36898: (case Term.strip_comb t of boehmes@36898: (q as Const (qn, _), [Abs (n, T, t')]) => boehmes@41059: if is_some (quantifier qn) then q $ Abs (n, as_tbool T, in_trig t') boehmes@36898: else as_term (in_term t) boehmes@41059: | (Const (c as (n as @{const_name distinct}, T)), [t']) => boehmes@41059: if B.is_builtin_fun ctxt c [t'] then boehmes@41059: Const (n, predT T) $ in_list T in_term t' boehmes@36898: else as_term (in_term t) boehmes@41059: | (Const (c as (n, T)), ts) => boehmes@41059: if B.is_builtin_conn ctxt c ts boehmes@41059: then Term.list_comb (Const c, map in_form ts) boehmes@41059: else if B.is_builtin_pred ctxt c ts boehmes@41059: then Term.list_comb (Const (n, predT T), map in_term ts) boehmes@36898: else as_term (in_term t) boehmes@36898: | _ => as_term (in_term t)) boehmes@36898: in boehmes@40161: map (apsnd (normalize ctxt)) #> (fn irules => boehmes@41059: ((rewrite_rules, (~1, term_bool) :: irules), boehmes@41059: term_bool_prop :: map (in_form o prop_of o snd) irules)) boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36898: (* translation from Isabelle terms into SMT intermediate terms *) boehmes@36898: boehmes@39298: val empty_context = (1, Typtab.empty, [], 1, Termtab.empty) boehmes@36898: boehmes@39298: fun make_sign header (_, typs, dtyps, _, terms) = { boehmes@36899: header = header, boehmes@39298: sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [], boehmes@39298: funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms [], boehmes@39435: dtyps = rev dtyps } boehmes@36898: boehmes@39298: fun make_recon (unfolds, assms) (_, typs, _, _, terms) = { boehmes@39298: typs = Symtab.make (map (apfst fst o swap) (Typtab.dest typs)), boehmes@39298: (*FIXME: don't drop the datatype information! *) boehmes@36898: terms = Symtab.make (map (fn (t, (n, _)) => (n, t)) (Termtab.dest terms)), boehmes@36898: unfolds = unfolds, boehmes@36899: assms = assms } boehmes@36898: boehmes@36898: fun string_of_index pre i = pre ^ string_of_int i boehmes@36898: boehmes@39298: fun new_typ sort_prefix proper T (Tidx, typs, dtyps, idx, terms) = boehmes@41059: let boehmes@41059: val s = string_of_index sort_prefix Tidx boehmes@41059: val U = revert_typ T boehmes@41059: in (s, (Tidx+1, Typtab.update (U, (s, proper)) typs, dtyps, idx, terms)) end boehmes@39298: boehmes@41059: fun lookup_typ (_, typs, _, _, _) = Typtab.lookup typs o revert_typ boehmes@36898: boehmes@39298: fun fresh_typ T f cx = boehmes@39298: (case lookup_typ cx T of boehmes@39298: SOME (s, _) => (s, cx) boehmes@39298: | NONE => f T cx) boehmes@39298: boehmes@39298: fun new_fun func_prefix t ss (Tidx, typs, dtyps, idx, terms) = boehmes@39298: let boehmes@39298: val f = string_of_index func_prefix idx boehmes@39298: val terms' = Termtab.update (revert_types t, (f, ss)) terms boehmes@39298: in (f, (Tidx, typs, dtyps, idx+1, terms')) end boehmes@39298: boehmes@39298: fun fresh_fun func_prefix t ss (cx as (_, _, _, _, terms)) = boehmes@41059: (case Termtab.lookup terms (revert_types t) of boehmes@36898: SOME (f, _) => (f, cx) boehmes@39298: | NONE => new_fun func_prefix t ss cx) boehmes@39298: boehmes@39483: fun mk_type (_, Tfs) (d as Datatype.DtTFree _) = the (AList.lookup (op =) Tfs d) boehmes@39483: | mk_type Ts (Datatype.DtType (n, ds)) = Type (n, map (mk_type Ts) ds) boehmes@39483: | mk_type (Tds, _) (Datatype.DtRec i) = nth Tds i boehmes@39298: boehmes@39483: fun mk_selector ctxt Ts T n (i, d) = boehmes@39483: (case Datatype_Selectors.lookup_selector ctxt (n, i+1) of boehmes@39483: NONE => raise Fail ("missing selector for datatype constructor " ^ quote n) boehmes@39483: | SOME m => mk_type Ts d |> (fn U => (Const (m, T --> U), U))) boehmes@39483: boehmes@39483: fun mk_constructor ctxt Ts T (n, args) = boehmes@39483: let val (sels, Us) = split_list (map_index (mk_selector ctxt Ts T n) args) boehmes@39483: in (Const (n, Us ---> T), sels) end boehmes@39298: boehmes@39535: fun lookup_datatype ctxt n Ts = boehmes@39535: if member (op =) [@{type_name bool}, @{type_name nat}] n then NONE boehmes@39535: else boehmes@39535: Datatype.get_info (ProofContext.theory_of ctxt) n boehmes@39535: |> Option.map (fn {descr, ...} => boehmes@39535: let boehmes@39535: val Tds = map (fn (_, (tn, _, _)) => Type (tn, Ts)) boehmes@39535: (sort (int_ord o pairself fst) descr) boehmes@39535: val Tfs = (case hd descr of (_, (_, tfs, _)) => tfs ~~ Ts) boehmes@39535: in boehmes@39535: descr |> map (fn (i, (_, _, cs)) => boehmes@39535: (nth Tds i, map (mk_constructor ctxt (Tds, Tfs) (nth Tds i)) cs)) boehmes@39535: end) boehmes@36898: boehmes@40161: fun relaxed irules = (([], irules), map (prop_of o snd) irules) boehmes@36898: boehmes@36899: fun with_context header f (ths, ts) = boehmes@36898: let val (us, context) = fold_map f ts empty_context boehmes@36899: in ((make_sign (header ts) context, us), make_recon ths context) end boehmes@36898: boehmes@36898: boehmes@41059: fun translate config ctxt comments = boehmes@36898: let boehmes@41059: val {prefixes, is_fol, header, has_datatypes, serialize} = config boehmes@36898: val {sort_prefix, func_prefix} = prefixes boehmes@39298: boehmes@39298: fun transT (T as TFree _) = fresh_typ T (new_typ sort_prefix true) boehmes@39298: | transT (T as TVar _) = (fn _ => raise TYPE ("smt_translate", [T], [])) boehmes@39298: | transT (T as Type (n, Ts)) = boehmes@41059: (case B.builtin_typ ctxt T of boehmes@39298: SOME n => pair n boehmes@39298: | NONE => fresh_typ T (fn _ => fn cx => boehmes@39298: if not has_datatypes then new_typ sort_prefix true T cx boehmes@39298: else boehmes@39298: (case lookup_datatype ctxt n Ts of boehmes@39298: NONE => new_typ sort_prefix true T cx boehmes@39298: | SOME dts => boehmes@39298: let val cx' = new_dtyps dts cx boehmes@39298: in (fst (the (lookup_typ cx' T)), cx') end))) boehmes@36898: boehmes@39298: and new_dtyps dts cx = boehmes@39298: let boehmes@39298: fun new_decl i t = boehmes@40663: let val (Ts, T) = U.dest_funT i (Term.fastype_of t) boehmes@39298: in boehmes@39298: fold_map transT Ts ##>> transT T ##>> boehmes@39298: new_fun func_prefix t NONE #>> swap boehmes@39298: end boehmes@39298: fun new_dtyp_decl (con, sels) = boehmes@39298: new_decl (length sels) con ##>> fold_map (new_decl 1) sels #>> boehmes@39298: (fn ((con', _), sels') => (con', map (apsnd snd) sels')) boehmes@39298: in boehmes@39298: cx boehmes@39298: |> fold_map (new_typ sort_prefix false o fst) dts boehmes@39298: ||>> fold_map (fold_map new_dtyp_decl o snd) dts boehmes@39298: |-> (fn (ss, decls) => fn (Tidx, typs, dtyps, idx, terms) => boehmes@39298: (Tidx, typs, (ss ~~ decls) :: dtyps, idx, terms)) boehmes@39298: end 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@36898: | NONE => raise TERM ("intermediate", [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@40681: | (h as Const (c as (@{const_name distinct}, T)), ts) => boehmes@41059: (case B.builtin_fun ctxt c ts of boehmes@36899: SOME (n, ts) => fold_map trans ts #>> app n boehmes@40681: | NONE => transs h T ts) boehmes@36898: | (h as Const (c as (_, T)), ts) => boehmes@41059: (case B.builtin_num ctxt t of boehmes@41059: SOME n => pair (SApp (n, [])) boehmes@36898: | NONE => boehmes@41059: (case B.builtin_fun ctxt c ts of boehmes@36899: SOME (n, ts') => fold_map trans ts' #>> app n boehmes@36898: | NONE => transs h T ts)) boehmes@36898: | (h as Free (_, T), ts) => transs h T ts boehmes@36898: | (Bound i, []) => pair (SVar i) blanchet@40697: | (Abs (_, _, t1 $ Bound 0), []) => blanchet@40697: if not (loose_bvar1 (t1, 0)) then trans t1 (* eta-reduce on the fly *) blanchet@40697: else raise TERM ("smt_translate", [t]) boehmes@39298: | _ => raise TERM ("smt_translate", [t])) boehmes@36898: 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@39298: fresh_fun func_prefix t (SOME Up) ##>> fold_map trans ts #>> SApp) boehmes@36898: end boehmes@36898: in boehmes@41059: (if is_fol then folify ctxt else relaxed) #> boehmes@36899: with_context (header ctxt) trans #>> uncurry (serialize comments) boehmes@36898: end boehmes@36898: boehmes@36898: end