src/Pure/Syntax/printer.ML
changeset 5691 3a6de95c09d0
parent 4699 fc5687450acc
child 6164 a0e9501d56f8
--- a/src/Pure/Syntax/printer.ML	Tue Oct 20 16:30:27 1998 +0200
+++ b/src/Pure/Syntax/printer.ML	Tue Oct 20 16:32:20 1998 +0200
@@ -17,29 +17,24 @@
 signature PRINTER =
 sig
   include PRINTER0
-  val term_to_ast: (string -> (bool -> typ -> term list -> term) option)
-    -> term -> Ast.ast
-  val typ_to_ast: (string -> (bool -> typ -> term list -> term) option)
-    -> typ -> Ast.ast
-  val sort_to_ast: (string -> (bool -> typ -> term list -> term) option)
-    -> sort -> Ast.ast
+  val term_to_ast: (string -> (bool -> typ -> term list -> term) list) -> term -> Ast.ast
+  val typ_to_ast: (string -> (bool -> typ -> term list -> term) list) -> typ -> Ast.ast
+  val sort_to_ast: (string -> (bool -> typ -> term list -> term) list) -> sort -> Ast.ast
   type prtabs
   val empty_prtabs: prtabs
   val extend_prtabs: prtabs -> string -> SynExt.xprod list -> prtabs
   val merge_prtabs: prtabs -> prtabs -> prtabs
   val pretty_term_ast: bool -> prtabs
-    -> (string -> (Ast.ast list -> Ast.ast) option)
+    -> (string -> (Ast.ast list -> Ast.ast) list)
     -> (string -> (string -> string * int) option) -> Ast.ast -> Pretty.T
   val pretty_typ_ast: bool -> prtabs
-    -> (string -> (Ast.ast list -> Ast.ast) option)
+    -> (string -> (Ast.ast list -> Ast.ast) list)
     -> (string -> (string -> string * int) option) -> Ast.ast -> Pretty.T
 end;
 
 structure Printer: PRINTER =
 struct
 
-open Lexicon Ast SynExt TypeExt SynTrans;
-
 
 (** options for printing **)
 
@@ -55,24 +50,27 @@
 
 (* apply print (ast) translation function *)
 
-fun apply_trans name a f args =
-  (f args handle
+fun apply_first [] x = raise Match
+  | apply_first (f :: fs) x = f x handle Match => apply_first fs x;
+
+fun apply_trans name a fs args =
+  (apply_first fs args handle
     Match => raise Match
   | exn => (writeln ("Error in " ^ name ^ " for " ^ quote a); raise exn));
 
-fun uncurry3 f (x, y, z) = f x y z;
+fun apply_typed x y fs = map (fn f => f x y) fs;
 
 
 (* simple_ast_of *)
 
-fun simple_ast_of (Const (c, _)) = Constant c
-  | simple_ast_of (Free (x, _)) = Variable x
-  | simple_ast_of (Var (xi, _)) = Variable (string_of_vname xi)
+fun simple_ast_of (Const (c, _)) = Ast.Constant c
+  | simple_ast_of (Free (x, _)) = Ast.Variable x
+  | simple_ast_of (Var (xi, _)) = Ast.Variable (Lexicon.string_of_vname xi)
   | simple_ast_of (t as _ $ _) =
       let val (f, args) = strip_comb t in
-        mk_appl (simple_ast_of f) (map simple_ast_of args)
+        Ast.mk_appl (simple_ast_of f) (map simple_ast_of args)
       end
-  | simple_ast_of (Bound i) = Variable ("B." ^ string_of_int i)
+  | simple_ast_of (Bound i) = Ast.Variable ("B." ^ string_of_int i)
   | simple_ast_of (Abs _) = sys_error "simple_ast_of: Abs";
 
 
@@ -81,6 +79,7 @@
 
 fun ast_of_termT trf tm =
   let
+    (* FIXME improve: lookup token classes *)
     fun ast_of (t as Const ("_class", _) $ Free _) = simple_ast_of t
       | ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of t
       | ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of t
@@ -88,20 +87,15 @@
       | ast_of (t as _ $ _) =
           (case strip_comb t of
             (Const (a, _), args) => trans a args
-          | (f, args) => Appl (map ast_of (f :: args)))
-      | ast_of t = raise TERM ("ast_of_termT: bad term encoding of type", [t])
+          | (f, args) => Ast.Appl (map ast_of (f :: args)))
+      | ast_of t = simple_ast_of t
     and trans a args =
-      (case trf a of
-        Some f => ast_of (apply_trans "print translation" a (uncurry3 f)
-          (false, dummyT, args))
-      | None => raise Match)
-          handle Match => mk_appl (Constant a) (map ast_of args);
-  in
-    ast_of tm
-  end;
+      ast_of (apply_trans "print translation" a (apply_typed false dummyT (trf a)) args)
+        handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args);
+  in ast_of tm end;
 
-fun sort_to_ast trf S = ast_of_termT trf (term_of_sort S);
-fun typ_to_ast trf T = ast_of_termT trf (term_of_typ (! show_sorts) T);
+fun sort_to_ast trf S = ast_of_termT trf (TypeExt.term_of_sort S);
+fun typ_to_ast trf T = ast_of_termT trf (TypeExt.term_of_typ (! show_sorts) T);
 
 
 
@@ -109,8 +103,8 @@
 
 fun mark_freevars (t $ u) = mark_freevars t $ mark_freevars u
   | mark_freevars (Abs (x, T, t)) = Abs (x, T, mark_freevars t)
-  | mark_freevars (t as Free _) = const "_free" $ t
-  | mark_freevars (t as Var _) = const "_var" $ t
+  | mark_freevars (t as Free _) = Lexicon.const "_free" $ t
+  | mark_freevars (t as Var _) =  Lexicon.const "_var" $ t
   | mark_freevars a = a;
 
 fun ast_of_term trf no_freeTs show_types show_sorts tm =
@@ -118,11 +112,11 @@
     fun prune_typs (t_seen as (Const _, _)) = t_seen
       | prune_typs (t as Free (x, ty), seen) =
           if ty = dummyT then (t, seen)
-          else if no_freeTs orelse t mem seen then (free x, seen)
+          else if no_freeTs orelse t mem seen then (Lexicon.free x, seen)
           else (t, t :: seen)
       | prune_typs (t as Var (xi, ty), seen) =
           if ty = dummyT then (t, seen)
-          else if no_freeTs orelse t mem seen then (var xi, seen)
+          else if no_freeTs orelse t mem seen then (Lexicon.var xi, seen)
           else (t, t :: seen)
       | prune_typs (t_seen as (Bound _, _)) = t_seen
       | prune_typs (Abs (x, ty, t), seen) =
@@ -138,32 +132,28 @@
 
     fun ast_of tm =
       (case strip_comb tm of
-        (t as Abs _, ts) => mk_appl (ast_of (abs_tr' t)) (map ast_of ts)
+        (t as Abs _, ts) => Ast.mk_appl (ast_of (SynTrans.abs_tr' t)) (map ast_of ts)
       | ((c as Const ("_free", _)), Free (x, T) :: ts) =>
-          mk_appl (constrain (c $ free x) T) (map ast_of ts)
+          Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts)
       | ((c as Const ("_bound", _)), Free (x, T) :: ts) =>
-          mk_appl (constrain (c $ free x) T) (map ast_of ts)
+          Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts)
       | ((c as Const ("_var", _)), Var (xi, T) :: ts) =>
-          mk_appl (constrain (c $ var xi) T) (map ast_of ts)
+          Ast.mk_appl (constrain (c $ Lexicon.var xi) T) (map ast_of ts)
       | (Const (c, T), ts) => trans c T ts
-      | (t, ts) => mk_appl (simple_ast_of t) (map ast_of ts))
+      | (t, ts) => Ast.mk_appl (simple_ast_of t) (map ast_of ts))
 
     and trans a T args =
-      (case trf a of
-        Some f =>
-          ast_of (apply_trans "print translation" a (uncurry3 f)
-            (show_sorts, T, args))
-      | None => raise Match)
-          handle Match => mk_appl (Constant a) (map ast_of args)
+      ast_of (apply_trans "print translation" a (apply_typed show_sorts T (trf a)) args)
+        handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args)
 
     and constrain t T =
       if show_types andalso T <> dummyT then
-        Appl [Constant constrainC, simple_ast_of t,
-          ast_of_termT trf (term_of_typ show_sorts T)]
+        Ast.Appl [Ast.Constant SynExt.constrainC, simple_ast_of t,
+          ast_of_termT trf (TypeExt.term_of_typ show_sorts T)]
       else simple_ast_of t
   in
     tm
-    |> prop_tr'
+    |> SynTrans.prop_tr'
     |> (if show_types then #1 o prune_typs o rpair [] else I)
     |> mark_freevars
     |> ast_of
@@ -202,32 +192,32 @@
 
 (* xprod_to_fmt *)
 
-fun xprod_to_fmt (XProd (_, _, "", _)) = None
-  | xprod_to_fmt (XProd (_, xsymbs, const, pri)) =
+fun xprod_to_fmt (SynExt.XProd (_, _, "", _)) = None
+  | xprod_to_fmt (SynExt.XProd (_, xsymbs, const, pri)) =
       let
         fun cons_str s (String s' :: syms) = String (s ^ s') :: syms
           | cons_str s syms = String s :: syms;
 
         fun arg (s, p) =
           (if s = "type" then TypArg else Arg)
-          (if is_terminal s then max_pri else p);
+          (if Lexicon.is_terminal s then SynExt.max_pri else p);
 
-        fun xsyms_to_syms (Delim s :: xsyms) =
+        fun xsyms_to_syms (SynExt.Delim s :: xsyms) =
               apfst (cons_str (Symbol.output s)) (xsyms_to_syms xsyms)
-          | xsyms_to_syms (Argument s_p :: xsyms) =
+          | xsyms_to_syms (SynExt.Argument s_p :: xsyms) =
               apfst (cons (arg s_p)) (xsyms_to_syms xsyms)
-          | xsyms_to_syms (Space s :: xsyms) =
+          | xsyms_to_syms (SynExt.Space s :: xsyms) =
               apfst (cons_str (Symbol.output s)) (xsyms_to_syms xsyms)
-          | xsyms_to_syms (Bg i :: xsyms) =
+          | xsyms_to_syms (SynExt.Bg i :: xsyms) =
               let
                 val (bsyms, xsyms') = xsyms_to_syms xsyms;
                 val (syms, xsyms'') = xsyms_to_syms xsyms';
               in
                 (Block (i, bsyms) :: syms, xsyms'')
               end
-          | xsyms_to_syms (Brk i :: xsyms) =
+          | xsyms_to_syms (SynExt.Brk i :: xsyms) =
               apfst (cons (Break i)) (xsyms_to_syms xsyms)
-          | xsyms_to_syms (En :: xsyms) = ([], xsyms)
+          | xsyms_to_syms (SynExt.En :: xsyms) = ([], xsyms)
           | xsyms_to_syms [] = ([], []);
 
         fun nargs (Arg _ :: syms) = nargs syms + 1
@@ -279,7 +269,7 @@
   let
     val trans = apply_trans "print ast translation";
 
-    fun token_trans c [Variable x] =
+    fun token_trans c [Ast.Variable x] =
           (case tokentrf c of
             None => None
           | Some f => Some (f x))
@@ -287,9 +277,9 @@
 
     (*default applications: prefix / postfix*)
     val appT =
-      if type_mode then tappl_ast_tr'
-      else if curried then applC_ast_tr'
-      else appl_ast_tr';
+      if type_mode then TypeExt.tappl_ast_tr'
+      else if curried then SynTrans.applC_ast_tr'
+      else SynTrans.appl_ast_tr';
 
     fun synT ([], args) = ([], args)
       | synT (Arg p :: symbs, t :: args) =
@@ -317,18 +307,18 @@
 
     and parT (pr, args, p, p': int) = #1 (synT
           (if p > p' orelse
-            (! show_brackets andalso p' <> max_pri andalso not (is_chain pr))
+            (! show_brackets andalso p' <> SynExt.max_pri andalso not (is_chain pr))
             then [Block (1, String "(" :: pr @ [String ")"])]
             else pr, args))
 
     and prefixT (_, a, [], _) = [Pretty.str a]
       | prefixT (c, _, args, p) =
-          if c = Constant "_appl" orelse c = Constant "_applC" then
+          if c = Ast.Constant "_appl" orelse c = Ast.Constant "_applC" then
             error "Syntax insufficient for printing prefix applications!"
           else astT (appT (c, args), p)
 
     and splitT 0 ([x], ys) = (x, ys)
-      | splitT 0 (rev_xs, ys) = (Appl (rev rev_xs), ys)
+      | splitT 0 (rev_xs, ys) = (Ast.Appl (rev rev_xs), ys)
       | splitT n (rev_xs, y :: ys) = splitT (n - 1) (y :: rev_xs, ys)
       | splitT _ _ = sys_error "splitT"
 
@@ -346,18 +336,16 @@
       in
         (case token_trans a args of     (*try token translation function*)
           Some (s, len) => [Pretty.strlen s len]
-        | None =>
-            (case trf a of              (*try print translation function*)
-              Some f => (astT (trans a f args, p) handle Match => prnt (get_fmts tabs a))
-            | None => prnt (get_fmts tabs a)))
+        | None =>			(*try print translation functions*)
+            astT (trans a (trf a) args, p) handle Match => prnt (get_fmts tabs a))
       end
 
-    and astT (c as Constant a, p) = combT (c, a, [], p)
-      | astT (Variable x, _) = [Pretty.str x]
-      | astT (Appl ((c as Constant a) :: (args as _ :: _)), p) =
+    and astT (c as Ast.Constant a, p) = combT (c, a, [], p)
+      | astT (Ast.Variable x, _) = [Pretty.str x]
+      | astT (Ast.Appl ((c as Ast.Constant a) :: (args as _ :: _)), p) =
           combT (c, a, args, p)
-      | astT (Appl (f :: (args as _ :: _)), p) = astT (appT (f, args), p)
-      | astT (ast as Appl _, _) = raise AST ("pretty: malformed ast", [ast]);
+      | astT (Ast.Appl (f :: (args as _ :: _)), p) = astT (appT (f, args), p)
+      | astT (ast as Ast.Appl _, _) = raise Ast.AST ("pretty: malformed ast", [ast]);
   in
     astT (ast0, p0)
   end;