src/Pure/Syntax/type_ext.ML
author clasohm
Fri, 22 Apr 1994 12:43:53 +0200
changeset 330 2fda15dd1e0f
parent 258 e540b7d4ecb1
child 347 cd41a57221d0
permissions -rw-r--r--
changed the way a grammar is generated to allow the new parser to work; also made a lot of changes in parser.ML and minor ones elsewhere

(*  Title:      Pure/Syntax/type_ext.ML
    ID:         $Id$
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen

The concrete syntax of types (used to bootstrap Pure).

TODO:
  term_of_typ: prune sorts
*)

signature TYPE_EXT0 =
sig
  val typ_of_term: (indexname -> sort) -> term -> typ
end;

signature TYPE_EXT =
sig
  include TYPE_EXT0
  structure SynExt: SYN_EXT
  local open SynExt SynExt.Ast in
    val term_of_typ: bool -> typ -> term
    val tappl_ast_tr': ast * ast list -> ast
    val type_ext: syn_ext
  end
end;

functor TypeExtFun(structure Lexicon: LEXICON and SynExt: SYN_EXT): TYPE_EXT =
struct

structure SynExt = SynExt;
open Lexicon SynExt SynExt.Ast;


(** typ_of_term **)

fun typ_of_term def_sort t =
  let
    fun sort_err (xi as (x, i)) =
      error ("Inconsistent sort constraints for type variable " ^
        quote (if i < 0 then x else string_of_vname xi));

    fun put_sort scs xi s =
      (case assoc (scs, xi) of
        None => (xi, s) :: scs
      | Some s' =>  if s = s' then scs else sort_err xi);

    fun insert x [] = [x: string]
      | insert x (lst as y :: ys) =
          if x > y then y :: insert x ys
          else if x = y then lst
          else x :: lst;

    fun classes (Const (c, _)) = [c]
      | classes (Free (c, _)) = [c]
      | classes (Const ("_classes", _) $ Const (c, _) $ cls) =
          insert c (classes cls)
      | classes (Const ("_classes", _) $ Free (c, _) $ cls) =
          insert c (classes cls)
      | classes tm = raise_term "typ_of_term: bad encoding of classes" [tm];

    fun sort (Const ("_emptysort", _)) = []
      | sort (Const (s, _)) = [s]
      | sort (Free (s, _)) = [s]
      | sort (Const ("_sort", _) $ cls) = classes cls
      | sort tm = raise_term "typ_of_term: bad encoding of sort" [tm];

    fun typ (Free (x, _), scs) =
          (if is_tid x then TFree (x, []) else Type (x, []), scs)
      | typ (Var (xi, _), scs) = (TVar (xi, []), scs)
      | typ (Const ("_ofsort", _) $ Free (x, _) $ st, scs) =
          (TFree (x, []), put_sort scs (x, ~1) (sort st))
      | typ (Const ("_ofsort", _) $ Var (xi, _) $ st, scs) =
          (TVar (xi, []), put_sort scs xi (sort st))
      | typ (Const (a, _), scs) = (Type (a, []), scs)
      | typ (tm as _ $ _, scs) =
          let
            val (t, ts) = strip_comb tm;
            val a =
              (case t of
                Const (x, _) => x
              | Free (x, _) => x
              | _ => raise_term "typ_of_term: bad type application" [tm]);
            val (tys, scs') = typs (ts, scs);
          in
            (Type (a, tys), scs')
          end
      | typ (tm, _) = raise_term "typ_of_term: bad encoding of typ" [tm]
    and typs (t :: ts, scs) =
          let
            val (ty, scs') = typ (t, scs);
            val (tys, scs'') = typs (ts, scs');
          in (ty :: tys, scs'') end
      | typs ([], scs) = ([], scs);


    val (ty, scs) = typ (t, []);

    fun get_sort xi =
      (case assoc (scs, xi) of
        None => def_sort xi
      | Some s => s);

    fun add_sorts (Type (a, tys)) = Type (a, map add_sorts tys)
      | add_sorts (TVar (xi, _)) = TVar (xi, get_sort xi)
      | add_sorts (TFree (x, _)) = TFree (x, get_sort (x, ~1));
  in
    add_sorts ty
  end;



(** term_of_typ **)

fun term_of_typ show_sorts ty =
  let
    fun const x = Const (x, dummyT);
    fun free x = Free (x, dummyT);
    fun var xi = Var (xi, dummyT);

    fun classes [] = raise Match
      | classes [c] = free c
      | classes (c :: cs) = const "_classes" $ free c $ classes cs;

    fun sort [] = const "_emptysort"
      | sort [s] = free s
      | sort ss = const "_sort" $ classes ss;

    fun of_sort t ss =
      if show_sorts then const "_ofsort" $ t $ sort ss else t;

    fun term_of (Type (a, tys)) = list_comb (const a, map term_of tys)
      | term_of (TFree (x, ss)) = of_sort (free x) ss
      | term_of (TVar (xi, ss)) = of_sort (var xi) ss;
  in
    term_of ty
  end;



(** the type syntax **)

(* parse ast translations *)

fun tappl_ast_tr (*"_tapp(l)"*) [types, f] =
      Appl (f :: unfold_ast "_types" types)
  | tappl_ast_tr (*"_tapp(l)"*) asts = raise_ast "tappl_ast_tr" asts;

fun bracket_ast_tr (*"_bracket"*) [dom, cod] =
      fold_ast_p "fun" (unfold_ast "_types" dom, cod)
  | bracket_ast_tr (*"_bracket"*) asts = raise_ast "bracket_ast_tr" asts;


(* print ast translations *)

fun tappl_ast_tr' (f, []) = raise_ast "tappl_ast_tr'" [f]
  | tappl_ast_tr' (f, [ty]) = Appl [Constant "_tapp", ty, f]
  | tappl_ast_tr' (f, tys) = Appl [Constant "_tappl", fold_ast "_types" tys, f];

fun fun_ast_tr' (*"fun"*) asts =
  (case unfold_ast_p "fun" (Appl (Constant "fun" :: asts)) of
    (dom as _ :: _ :: _, cod)
      => Appl [Constant "_bracket", fold_ast "_types" dom, cod]
  | _ => raise Match);


(* type_ext *)

val sortT = Type ("sort", []);
val classesT = Type ("classes", []);
val typesT = Type ("types", []);

val type_ext = syn_ext
  [logic, "type"] [logic, "type"]
  [Mfix ("_",           tidT --> typeT,              "", [], max_pri),
   Mfix ("_",           tvarT --> typeT,               "", [], max_pri),
   Mfix ("_",           idT --> typeT,                 "", [], max_pri),
   Mfix ("_::_",        [tidT, sortT] ---> typeT,      "_ofsort", [max_pri, 0], max_pri),
   Mfix ("_::_",        [tvarT, sortT] ---> typeT,     "_ofsort", [max_pri, 0], max_pri),
   Mfix ("_",           idT --> sortT,                 "", [], max_pri),
   Mfix ("{}",          sortT,                         "_emptysort", [], max_pri),
   Mfix ("{_}",         classesT --> sortT,            "_sort", [], max_pri),
   Mfix ("_",           idT --> classesT,              "", [], max_pri),
   Mfix ("_,_",         [idT, classesT] ---> classesT, "_classes", [], max_pri),
   Mfix ("_ _",         [typeT, idT] ---> typeT,       "_tapp", [max_pri, 0], max_pri),
   Mfix ("((1'(_,/ _'))_)", [typeT, typesT, idT] ---> typeT, "_tappl", [], max_pri),
   Mfix ("_",           typeT --> typesT,              "", [], max_pri),
   Mfix ("_,/ _",       [typeT, typesT] ---> typesT,   "_types", [], max_pri),
   Mfix ("(_/ => _)",   [typeT, typeT] ---> typeT,     "fun", [1, 0], 0),
   Mfix ("([_]/ => _)", [typesT, typeT] ---> typeT,    "_bracket", [0, 0], 0)]
  []
  ([("_tapp", tappl_ast_tr), ("_tappl", tappl_ast_tr), ("_bracket", bracket_ast_tr)],
   [],
   [],
   [("fun", fun_ast_tr')])
  ([], []);


end;