(* Title: Pure/Syntax/syn_trans.ML
ID: $Id$
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen
Syntax translation functions.
*)
signature SYN_TRANS0 =
sig
val eta_contract: bool ref
val mk_binder_tr: string * string -> string * (term list -> term)
val mk_binder_tr': string * string -> string * (term list -> term)
val dependent_tr': string * string -> term list -> term
val antiquote_tr: string -> term -> term
val quote_tr: string -> term -> term
val quote_antiquote_tr: string -> string -> string -> string * (term list -> term)
val antiquote_tr': string -> term -> term
val quote_tr': string -> term -> term
val quote_antiquote_tr': string -> string -> string -> string * (term list -> term)
val mark_bound: string -> term
val mark_boundT: string * typ -> term
val variant_abs': string * typ * term -> string * term
end;
signature SYN_TRANS1 =
sig
include SYN_TRANS0
val constrainAbsC: string
val pure_trfuns:
(string * (Ast.ast list -> Ast.ast)) list *
(string * (term list -> term)) list *
(string * (term list -> term)) list *
(string * (Ast.ast list -> Ast.ast)) list
val pure_trfunsT: (string * (bool -> typ -> term list -> term)) list
end;
signature SYN_TRANS =
sig
include SYN_TRANS1
val abs_tr': term -> term
val prop_tr': term -> term
val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
val pt_to_ast: (string -> (Ast.ast list -> Ast.ast) option) -> Parser.parsetree -> Ast.ast
val ast_to_term: (string -> (term list -> term) option) -> Ast.ast -> term
end;
structure SynTrans: SYN_TRANS =
struct
(** parse (ast) translations **)
(* application *)
fun appl_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_args" args)
| appl_ast_tr asts = raise Ast.AST ("appl_ast_tr", asts);
fun applC_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_cargs" args)
| applC_ast_tr asts = raise Ast.AST ("applC_ast_tr", asts);
(* abstraction *)
fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Ast.Appl [Ast.Constant SynExt.constrainC, x, ty]
| idtyp_ast_tr (*"_idtyp"*) asts = raise Ast.AST ("idtyp_ast_tr", asts);
fun lambda_ast_tr (*"_lambda"*) [pats, body] =
Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body)
| lambda_ast_tr (*"_lambda"*) asts = raise Ast.AST ("lambda_ast_tr", asts);
val constrainAbsC = "_constrainAbs";
fun abs_tr (*"_abs"*) [Free (x, T), body] = Term.absfree (x, T, body)
| abs_tr (*"_abs"*) (ts as [Const (c, _) $ Free (x, T) $ tT, body]) =
if c = SynExt.constrainC
then Lexicon.const constrainAbsC $ Term.absfree (x, T, body) $ tT
else raise TERM ("abs_tr", ts)
| abs_tr (*"_abs"*) ts = raise TERM ("abs_tr", ts);
(* nondependent abstraction *)
fun k_tr (*"_K"*) [t] = Abs ("uu", dummyT, Term.incr_boundvars 1 t)
| k_tr (*"_K"*) ts = raise TERM ("k_tr", ts);
(* binder *)
fun mk_binder_tr (sy, name) =
let
fun tr (Free (x, T), t) = Lexicon.const name $ Term.absfree (x, T, t)
| tr (Const ("_idts", _) $ idt $ idts, t) = tr (idt, tr (idts, t))
| tr (t1 as Const (c, _) $ Free (x, T) $ tT, t) =
if c = SynExt.constrainC then
Lexicon.const name $ (Lexicon.const constrainAbsC $ Term.absfree (x, T, t) $ tT)
else raise TERM ("binder_tr", [t1, t])
| tr (t1, t2) = raise TERM ("binder_tr", [t1, t2]);
fun binder_tr (*sy*) [idts, body] = tr (idts, body)
| binder_tr (*sy*) ts = raise TERM ("binder_tr", ts);
in
(sy, binder_tr)
end;
(* meta propositions *)
fun aprop_tr (*"_aprop"*) [t] = Lexicon.const SynExt.constrainC $ t $ Lexicon.const "prop"
| aprop_tr (*"_aprop"*) ts = raise TERM ("aprop_tr", ts);
fun ofclass_tr (*"_ofclass"*) [ty, cls] =
cls $ (Lexicon.const SynExt.constrainC $ Lexicon.const "TYPE" $
(Lexicon.const "itself" $ ty))
| ofclass_tr (*"_ofclass"*) ts = raise TERM ("ofclass_tr", ts);
(* meta implication *)
fun bigimpl_ast_tr (*"_bigimpl"*) [asms, concl] =
Ast.fold_ast_p "==>" (Ast.unfold_ast "_asms" asms, concl)
| bigimpl_ast_tr (*"_bigimpl"*) asts = raise Ast.AST ("bigimpl_ast_tr", asts);
(* type reflection *)
fun type_tr (*"_TYPE"*) [ty] =
Lexicon.const SynExt.constrainC $ Lexicon.const "TYPE" $ (Lexicon.const "itself" $ ty)
| type_tr (*"_TYPE"*) ts = raise TERM ("type_tr", ts);
(* dddot *)
fun dddot_tr (*"_DDDOT"*) ts = Term.list_comb (Lexicon.var SynExt.dddot_indexname, ts);
(* quote / antiquote *)
fun antiquote_tr name =
let
fun tr i ((t as Const (c, _)) $ u) =
if c = name then tr i u $ Bound i
else tr i t $ tr i u
| tr i (t $ u) = tr i t $ tr i u
| tr i (Abs (x, T, t)) = Abs (x, T, tr (i + 1) t)
| tr _ a = a;
in tr 0 end;
fun quote_tr name t = Abs ("s", dummyT, antiquote_tr name (Term.incr_boundvars 1 t));
fun quote_antiquote_tr quoteN antiquoteN name =
let
fun tr [t] = Lexicon.const name $ quote_tr antiquoteN t
| tr ts = raise TERM ("quote_tr", ts);
in (quoteN, tr) end;
(** print (ast) translations **)
(* application *)
fun appl_ast_tr' (f, []) = raise Ast.AST ("appl_ast_tr'", [f])
| appl_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_appl", f, Ast.fold_ast "_args" args];
fun applC_ast_tr' (f, []) = raise Ast.AST ("applC_ast_tr'", [f])
| applC_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_applC", f, Ast.fold_ast "_cargs" args];
(* abstraction *)
fun mark_boundT x_T = Lexicon.const "_bound" $ Free x_T;
fun mark_bound x = mark_boundT (x, dummyT);
fun strip_abss vars_of body_of tm =
let
val vars = vars_of tm;
val body = body_of tm;
val rev_new_vars = rename_wrt_term body vars;
in
(map mark_boundT (rev rev_new_vars),
subst_bounds (map (mark_bound o #1) rev_new_vars, body))
end;
(*do (partial) eta-contraction before printing*)
val eta_contract = ref true;
fun eta_contr tm =
let
fun is_aprop (Const ("_aprop", _)) = true
| is_aprop _ = false;
fun eta_abs (Abs (a, T, t)) =
(case eta_abs t of
t' as f $ u =>
(case eta_abs u of
Bound 0 =>
if Term.loose_bvar1 (f, 0) orelse is_aprop f then Abs (a, T, t')
else incr_boundvars ~1 f
| _ => Abs (a, T, t'))
| t' => Abs (a, T, t'))
| eta_abs t = t;
in
if ! eta_contract then eta_abs tm else tm
end;
fun abs_tr' tm =
foldr (fn (x, t) => Lexicon.const "_abs" $ x $ t)
(strip_abss strip_abs_vars strip_abs_body (eta_contr tm));
fun abs_ast_tr' (*"_abs"*) asts =
(case Ast.unfold_ast_p "_abs" (Ast.Appl (Ast.Constant "_abs" :: asts)) of
([], _) => raise Ast.AST ("abs_ast_tr'", asts)
| (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]);
(* binder *)
fun mk_binder_tr' (name, sy) =
let
fun mk_idts [] = raise Match (*abort translation*)
| mk_idts [idt] = idt
| mk_idts (idt :: idts) = Lexicon.const "_idts" $ idt $ mk_idts idts;
fun tr' t =
let
val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t;
in Lexicon.const sy $ mk_idts xs $ bd end;
fun binder_tr' (*name*) (t :: ts) = Term.list_comb (tr' (Lexicon.const name $ t), ts)
| binder_tr' (*name*) [] = raise Match;
in
(name, binder_tr')
end;
(* idtyp constraints *)
fun idtyp_ast_tr' a [Ast.Appl [Ast.Constant c, x, ty], xs] =
if c = SynExt.constrainC then
Ast.Appl [ Ast.Constant a, Ast.Appl [Ast.Constant "_idtyp", x, ty], xs]
else raise Match
| idtyp_ast_tr' _ _ = raise Match;
(* meta propositions *)
fun prop_tr' tm =
let
fun aprop t = Lexicon.const "_aprop" $ t;
fun is_prop Ts t =
fastype_of1 (Ts, t) = propT handle TERM _ => false;
fun tr' _ (t as Const _) = t
| tr' _ (t as Free (x, T)) =
if T = propT then aprop (Lexicon.free x) else t
| tr' _ (t as Var (xi, T)) =
if T = propT then aprop (Lexicon.var xi) else t
| tr' Ts (t as Bound _) =
if is_prop Ts t then aprop t else t
| tr' Ts (Abs (x, T, t)) = Abs (x, T, tr' (T :: Ts) t)
| tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) =
if is_prop Ts t then Const ("_mk_ofclass", T) $ tr' Ts t1
else tr' Ts t1 $ tr' Ts t2
| tr' Ts (t as t1 $ t2) =
(if is_Const (Term.head_of t) orelse not (is_prop Ts t)
then I else aprop) (tr' Ts t1 $ tr' Ts t2);
in
tr' [] tm
end;
fun mk_ofclass_tr' show_sorts (*"_mk_ofclass"*) T [t] =
Lexicon.const "_ofclass" $ TypeExt.term_of_typ show_sorts T $ t
| mk_ofclass_tr' _ (*"_mk_ofclass"*) T ts = raise TYPE ("mk_ofclass_tr'", [T], ts);
(* meta implication *)
fun impl_ast_tr' (*"==>"*) asts =
(case Ast.unfold_ast_p "==>" (Ast.Appl (Ast.Constant "==>" :: asts)) of
(asms as _ :: _ :: _, concl)
=> Ast.Appl [Ast.Constant "_bigimpl", Ast.fold_ast "_asms" asms, concl]
| _ => raise Match);
(* type reflection *)
fun type_tr' show_sorts (*"TYPE"*) (Type ("itself", [T])) ts =
Term.list_comb (Lexicon.const "_TYPE" $ TypeExt.term_of_typ show_sorts T, ts)
| type_tr' _ _ _ = raise Match;
(* dependent / nondependent quantifiers *)
fun variant_abs' (x, T, B) =
let val x' = variant (add_term_names (B, [])) x in
(x', subst_bound (mark_boundT (x', T), B))
end;
fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) =
if Term.loose_bvar1 (B, 0) then
let val (x', B') = variant_abs' (x, dummyT, B);
in Term.list_comb (Lexicon.const q $ mark_boundT (x', T) $ A $ B', ts) end
else Term.list_comb (Lexicon.const r $ A $ B, ts)
| dependent_tr' _ _ = raise Match;
(* quote / antiquote *)
fun antiquote_tr' name =
let
fun tr' i (t $ u) =
if u = Bound i then Lexicon.const name $ tr' i t
else tr' i t $ tr' i u
| tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t)
| tr' i a = if a = Bound i then raise Match else a;
in tr' 0 end;
fun quote_tr' name (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_tr' name t)
| quote_tr' _ _ = raise Match;
fun quote_antiquote_tr' quoteN antiquoteN name =
let
fun tr' (t :: ts) = Term.list_comb (Lexicon.const quoteN $ quote_tr' antiquoteN t, ts)
| tr' _ = raise Match;
in (name, tr') end;
(** pure_trfuns **)
val pure_trfuns =
([("_appl", appl_ast_tr), ("_applC", applC_ast_tr),
("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr),
("_bigimpl", bigimpl_ast_tr)],
[("_abs", abs_tr), ("_aprop", aprop_tr), ("_ofclass", ofclass_tr),
("_TYPE", type_tr), ("_DDDOT", dddot_tr), ("_K", k_tr)],
[]: (string * (term list -> term)) list,
[("_abs", abs_ast_tr'), ("_idts", idtyp_ast_tr' "_idts"),
("_pttrns", idtyp_ast_tr' "_pttrns"), ("==>", impl_ast_tr')]);
val pure_trfunsT =
[("_mk_ofclass", mk_ofclass_tr'), ("TYPE", type_tr')];
(** pt_to_ast **)
fun pt_to_ast trf pt =
let
fun trans a args =
(case trf a of
None => Ast.mk_appl (Ast.Constant a) args
| Some f => f args handle exn
=> (writeln ("Error in parse ast translation for " ^ quote a);
raise exn));
(*translate pt bottom-up*)
fun ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts)
| ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok);
in
ast_of pt
end;
(** ast_to_term **)
fun ast_to_term trf ast =
let
fun trans a args =
(case trf a of
None => Term.list_comb (Lexicon.const a, args)
| Some f => f args handle exn
=> (writeln ("Error in parse translation for " ^ quote a);
raise exn));
fun term_of (Ast.Constant a) = trans a []
| term_of (Ast.Variable x) = Lexicon.read_var x
| term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) =
trans a (map term_of asts)
| term_of (Ast.Appl (ast :: (asts as _ :: _))) =
Term.list_comb (term_of ast, map term_of asts)
| term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]);
in
term_of ast
end;
end;