src/Pure/Syntax/printer.ML
author wenzelm
Wed, 03 Mar 2010 00:28:22 +0100
changeset 35429 afa8cf9e63d8
parent 35262 9ea4445d2ccf
child 35435 e6c03f397eb8
permissions -rw-r--r--
authentic syntax for classes and type constructors; read/intern formal entities just after raw parsing, extern just before final pretty printing; discontinued _class token translation; moved Local_Syntax.extern_term to Syntax/printer.ML; misc tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
     1
(*  Title:      Pure/Syntax/printer.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
     3
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
     4
Pretty printing of asts, terms, types and print (ast) translation.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
signature PRINTER0 =
2384
d360b395766e removed chartrans_of;
wenzelm
parents: 2365
diff changeset
     8
sig
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
     9
  val show_brackets: bool Unsynchronized.ref
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    10
  val show_sorts: bool Unsynchronized.ref
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    11
  val show_types: bool Unsynchronized.ref
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    12
  val show_no_free_types: bool Unsynchronized.ref
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    13
  val show_all_types: bool Unsynchronized.ref
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    14
  val show_structs: bool Unsynchronized.ref
14837
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
    15
  val pp_show_brackets: Pretty.pp -> Pretty.pp
2384
d360b395766e removed chartrans_of;
wenzelm
parents: 2365
diff changeset
    16
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
signature PRINTER =
2384
d360b395766e removed chartrans_of;
wenzelm
parents: 2365
diff changeset
    19
sig
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
  include PRINTER0
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    21
  val sort_to_ast: Proof.context ->
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    22
    (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> sort -> Ast.ast
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21748
diff changeset
    23
  val typ_to_ast: Proof.context ->
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21748
diff changeset
    24
    (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> typ -> Ast.ast
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    25
  val term_to_ast: {structs: string list, fixes: string list} -> string list -> Proof.context ->
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    26
    (string -> (Proof.context -> bool -> typ -> term list -> term) list) -> term -> Ast.ast
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
    27
  type prtabs
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
    28
  val empty_prtabs: prtabs
25393
0856e0141062 replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents: 25386
diff changeset
    29
  val update_prtabs: string -> SynExt.xprod list -> prtabs -> prtabs
15753
eb014dfc57ee tuned extend_prtabs;
wenzelm
parents: 15574
diff changeset
    30
  val remove_prtabs: string -> SynExt.xprod list -> prtabs -> prtabs
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
    31
  val merge_prtabs: prtabs -> prtabs -> prtabs
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    32
  val pretty_term_ast: {extern_class: string -> xstring, extern_type: string -> xstring,
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    33
      extern_const: string -> xstring} -> Proof.context -> bool -> prtabs ->
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    34
    (string -> (Proof.context -> Ast.ast list -> Ast.ast) list) ->
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    35
    (string -> (Proof.context -> string -> Pretty.T) option) -> Ast.ast -> Pretty.T list
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    36
  val pretty_typ_ast: {extern_class: string -> xstring, extern_type: string -> xstring} ->
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    37
    Proof.context -> bool -> prtabs ->
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    38
    (string -> (Proof.context -> Ast.ast list -> Ast.ast) list) ->
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    39
    (string -> (Proof.context -> string -> Pretty.T) option) -> Ast.ast -> Pretty.T list
2384
d360b395766e removed chartrans_of;
wenzelm
parents: 2365
diff changeset
    40
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
2365
38295260a740 added chartrans_of;
wenzelm
parents: 2229
diff changeset
    42
structure Printer: PRINTER =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
struct
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
    44
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
    45
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
(** options for printing **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    48
val show_types = Unsynchronized.ref false;
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    49
val show_sorts = Unsynchronized.ref false;
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    50
val show_brackets = Unsynchronized.ref false;
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    51
val show_no_free_types = Unsynchronized.ref false;
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 29565
diff changeset
    52
val show_all_types = Unsynchronized.ref false;
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    53
val show_structs = Unsynchronized.ref false;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
32966
5b21661fe618 indicate CRITICAL nature of various setmp combinators;
wenzelm
parents: 32738
diff changeset
    55
fun pp_show_brackets pp = Pretty.pp (setmp_CRITICAL show_brackets true (Pretty.term pp),
14975
2736b0984253 tuned pp;
wenzelm
parents: 14837
diff changeset
    56
  Pretty.typ pp, Pretty.sort pp, Pretty.classrel pp, Pretty.arity pp);
14837
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
    57
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
    58
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    60
(** misc utils **)
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    61
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    62
(* apply print (ast) translation function *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
23937
66e1f24d655d eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents: 23660
diff changeset
    64
fun apply_trans ctxt fns args =
16611
edb368e2878f proper treatment of advanced trfuns: pass thy argument;
wenzelm
parents: 15973
diff changeset
    65
  let
edb368e2878f proper treatment of advanced trfuns: pass thy argument;
wenzelm
parents: 15973
diff changeset
    66
    fun app_first [] = raise Match
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21748
diff changeset
    67
      | app_first (f :: fs) = f ctxt args handle Match => app_first fs;
23937
66e1f24d655d eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents: 23660
diff changeset
    68
  in app_first fns end;
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
    69
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21748
diff changeset
    70
fun apply_typed x y fs = map (fn f => fn ctxt => f ctxt x y) fs;
4147
e57d03a5fc73 print translation: added show_sorts argument;
wenzelm
parents: 3816
diff changeset
    71
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    73
(* simple_ast_of *)
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    74
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
    75
fun simple_ast_of (Const (c, _)) = Ast.Constant c
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
    76
  | simple_ast_of (Free (x, _)) = Ast.Variable x
14783
e7f7ed4c06f2 string_of_vname moved to term.ML;
wenzelm
parents: 14696
diff changeset
    77
  | simple_ast_of (Var (xi, _)) = Ast.Variable (Term.string_of_vname xi)
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    78
  | simple_ast_of (t as _ $ _) =
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    79
      let val (f, args) = strip_comb t in
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
    80
        Ast.mk_appl (simple_ast_of f) (map simple_ast_of args)
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    81
      end
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
    82
  | simple_ast_of (Bound i) = Ast.Variable ("B." ^ string_of_int i)
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    83
  | simple_ast_of (Abs _) = sys_error "simple_ast_of: Abs";
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    84
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    85
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    86
3776
38f8ec304b95 added sort_to_ast;
wenzelm
parents: 2913
diff changeset
    87
(** sort_to_ast, typ_to_ast **)
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    88
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21748
diff changeset
    89
fun ast_of_termT ctxt trf tm =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
  let
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    91
    fun ast_of (t as Const ("_tfree", _) $ Free _) = simple_ast_of t
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    92
      | ast_of (t as Const ("_tvar", _) $ Var _) = simple_ast_of t
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    93
      | ast_of (Const (a, _)) = trans a []
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    94
      | ast_of (t as _ $ _) =
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    95
          (case strip_comb t of
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    96
            (Const (a, _), args) => trans a args
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
    97
          | (f, args) => Ast.Appl (map ast_of (f :: args)))
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
    98
      | ast_of t = simple_ast_of t
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
    99
    and trans a args =
23937
66e1f24d655d eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents: 23660
diff changeset
   100
      ast_of (apply_trans ctxt (apply_typed false dummyT (trf a)) args)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   101
        handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args);
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   102
  in ast_of tm end;
617
94436ad4b60d added flag show_no_free_types: bool ref;
wenzelm
parents: 554
diff changeset
   103
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21748
diff changeset
   104
fun sort_to_ast ctxt trf S = ast_of_termT ctxt trf (TypeExt.term_of_sort S);
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21748
diff changeset
   105
fun typ_to_ast ctxt trf T = ast_of_termT ctxt trf (TypeExt.term_of_typ (! show_sorts) T);
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   106
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   107
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   108
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   109
(** term_to_ast **)
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   110
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   111
fun ast_of_term idents consts ctxt trf
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   112
    show_all_types no_freeTs show_types show_sorts show_structs tm =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   113
  let
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   114
    val {structs, fixes} = idents;
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   115
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   116
    fun mark_atoms ((t as Const (c, T)) $ u) =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   117
          if member (op =) consts c then (t $ u)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   118
          else Const (Lexicon.mark_const c, T) $ mark_atoms u
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   119
      | mark_atoms (t $ u) = mark_atoms t $ mark_atoms u
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   120
      | mark_atoms (Abs (x, T, t)) = Abs (x, T, mark_atoms t)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   121
      | mark_atoms (Const (c, T)) = Const (Lexicon.mark_const c, T)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   122
      | mark_atoms (t as Free (x, T)) =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   123
          let val i = find_index (fn s => s = x) structs + 1 in
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   124
            if i = 0 andalso member (op =) fixes x then
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   125
              Term.Const (Lexicon.mark_fixed x, T)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   126
            else if i = 1 andalso not show_structs then
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   127
              Lexicon.const "_struct" $ Lexicon.const "_indexdefault"
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   128
            else Lexicon.const "_free" $ t
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   129
          end
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   130
      | mark_atoms (t as Var (xi, T)) =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   131
          if xi = SynExt.dddot_indexname then Const ("_DDDOT", T)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   132
          else Lexicon.const "_var" $ t
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   133
      | mark_atoms a = a;
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   134
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
    fun prune_typs (t_seen as (Const _, _)) = t_seen
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   136
      | prune_typs (t as Free (x, ty), seen) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
          if ty = dummyT then (t, seen)
18139
wenzelm
parents: 17412
diff changeset
   138
          else if no_freeTs orelse member (op aconv) seen t then (Lexicon.free x, seen)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
          else (t, t :: seen)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
      | prune_typs (t as Var (xi, ty), seen) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   141
          if ty = dummyT then (t, seen)
18139
wenzelm
parents: 17412
diff changeset
   142
          else if no_freeTs orelse member (op aconv) seen t then (Lexicon.var xi, seen)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
          else (t, t :: seen)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
      | prune_typs (t_seen as (Bound _, _)) = t_seen
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   145
      | prune_typs (Abs (x, ty, t), seen) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
          let val (t', seen') = prune_typs (t, seen);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   147
          in (Abs (x, ty, t'), seen') end
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
      | prune_typs (t1 $ t2, seen) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   149
          let
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   150
            val (t1', seen') = prune_typs (t1, seen);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   151
            val (t2', seen'') = prune_typs (t2, seen');
6767
99797f2652d1 print "..." variable;
wenzelm
parents: 6322
diff changeset
   152
          in (t1' $ t2', seen'') end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   153
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   154
    fun ast_of tm =
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   155
      (case strip_comb tm of
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   156
        (t as Abs _, ts) => Ast.mk_appl (ast_of (SynTrans.abs_tr' t)) (map ast_of ts)
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   157
      | ((c as Const ("_free", _)), Free (x, T) :: ts) =>
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   158
          Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts)
21748
7df0f4e08dde support printing of idtdummy;
wenzelm
parents: 19482
diff changeset
   159
      | ((c as Const ("_var", _)), Var (xi, T) :: ts) =>
7df0f4e08dde support printing of idtdummy;
wenzelm
parents: 19482
diff changeset
   160
          Ast.mk_appl (constrain (c $ Lexicon.var xi) T) (map ast_of ts)
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   161
      | ((c as Const ("_bound", _)), Free (x, T) :: ts) =>
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   162
          Ast.mk_appl (constrain (c $ Lexicon.free x) T) (map ast_of ts)
21748
7df0f4e08dde support printing of idtdummy;
wenzelm
parents: 19482
diff changeset
   163
      | (Const ("_idtdummy", T), ts) =>
7df0f4e08dde support printing of idtdummy;
wenzelm
parents: 19482
diff changeset
   164
          Ast.mk_appl (constrain (Lexicon.const "_idtdummy") T) (map ast_of ts)
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   165
      | (const as Const (c, T), ts) =>
14696
wenzelm
parents: 14176
diff changeset
   166
          if show_all_types
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   167
          then Ast.mk_appl (constrain const T) (map ast_of ts)
14696
wenzelm
parents: 14176
diff changeset
   168
          else trans c T ts
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   169
      | (t, ts) => Ast.mk_appl (simple_ast_of t) (map ast_of ts))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
2384
d360b395766e removed chartrans_of;
wenzelm
parents: 2365
diff changeset
   171
    and trans a T args =
23937
66e1f24d655d eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents: 23660
diff changeset
   172
      ast_of (apply_trans ctxt (apply_typed show_sorts T (trf a)) args)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   173
        handle Match => Ast.mk_appl (Ast.Constant a) (map ast_of args)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   174
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   175
    and constrain t T =
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   176
      if show_types andalso T <> dummyT then
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   177
        Ast.Appl [Ast.Constant SynExt.constrainC, simple_ast_of t,
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21748
diff changeset
   178
          ast_of_termT ctxt trf (TypeExt.term_of_typ show_sorts T)]
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   179
      else simple_ast_of t;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   180
  in
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   181
    tm
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   182
    |> SynTrans.prop_tr'
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   183
    |> show_types ? (#1 o prune_typs o rpair [])
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   184
    |> mark_atoms
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   185
    |> ast_of
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   186
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   188
fun term_to_ast idents consts ctxt trf tm =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   189
  ast_of_term idents consts ctxt trf (! show_all_types) (! show_no_free_types)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   190
    (! show_types orelse ! show_sorts orelse ! show_all_types) (! show_sorts) (! show_structs) tm;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   191
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   192
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   194
(** type prtabs **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   195
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   196
datatype symb =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
  Arg of int |
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   198
  TypArg of int |
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
  String of string |
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   200
  Space of string |
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   201
  Break of int |
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   202
  Block of int * symb list;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   203
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   204
type prtabs = (string * ((symb list * int * int) list) Symtab.table) list;
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   205
18957
8c3abd63bce3 TableFun: renamed xxx_multi to xxx_list;
wenzelm
parents: 18857
diff changeset
   206
fun mode_tab prtabs mode = the_default Symtab.empty (AList.lookup (op =) prtabs mode);
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19374
diff changeset
   207
fun mode_tabs prtabs modes = map_filter (AList.lookup (op =) prtabs) (modes @ [""]);
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   208
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   209
3816
7e1b695bcc5e changed preference order of prtab entries;
wenzelm
parents: 3776
diff changeset
   210
(* xprod_to_fmt *)
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   211
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   212
fun xprod_to_fmt (SynExt.XProd (_, _, "", _)) = NONE
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   213
  | xprod_to_fmt (SynExt.XProd (_, xsymbs, const, pri)) =
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   214
      let
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   215
        fun arg (s, p) =
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   216
          (if s = "type" then TypArg else Arg)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   217
          (if Lexicon.is_terminal s then SynExt.max_pri else p);
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   218
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   219
        fun xsyms_to_syms (SynExt.Delim s :: xsyms) =
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   220
              apfst (cons (String s)) (xsyms_to_syms xsyms)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   221
          | xsyms_to_syms (SynExt.Argument s_p :: xsyms) =
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   222
              apfst (cons (arg s_p)) (xsyms_to_syms xsyms)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   223
          | xsyms_to_syms (SynExt.Space s :: xsyms) =
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   224
              apfst (cons (Space s)) (xsyms_to_syms xsyms)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   225
          | xsyms_to_syms (SynExt.Bg i :: xsyms) =
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   226
              let
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   227
                val (bsyms, xsyms') = xsyms_to_syms xsyms;
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   228
                val (syms, xsyms'') = xsyms_to_syms xsyms';
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   229
              in
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   230
                (Block (i, bsyms) :: syms, xsyms'')
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   231
              end
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   232
          | xsyms_to_syms (SynExt.Brk i :: xsyms) =
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   233
              apfst (cons (Break i)) (xsyms_to_syms xsyms)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   234
          | xsyms_to_syms (SynExt.En :: xsyms) = ([], xsyms)
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   235
          | xsyms_to_syms [] = ([], []);
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   236
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   237
        fun nargs (Arg _ :: syms) = nargs syms + 1
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   238
          | nargs (TypArg _ :: syms) = nargs syms + 1
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   239
          | nargs (String _ :: syms) = nargs syms
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   240
          | nargs (Space _ :: syms) = nargs syms
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   241
          | nargs (Break _ :: syms) = nargs syms
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   242
          | nargs (Block (_, bsyms) :: syms) = nargs syms + nargs bsyms
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   243
          | nargs [] = 0;
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   244
      in
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   245
        (case xsyms_to_syms xsymbs of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   246
          (symbs, []) => SOME (const, (symbs, nargs symbs, pri))
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   247
        | _ => sys_error "xprod_to_fmt: unbalanced blocks")
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   248
      end;
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   249
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   250
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   251
(* empty, extend, merge prtabs *)
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   252
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   253
val empty_prtabs = [];
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   254
25393
0856e0141062 replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents: 25386
diff changeset
   255
fun update_prtabs mode xprods prtabs =
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   256
  let
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19374
diff changeset
   257
    val fmts = map_filter xprod_to_fmt xprods;
25393
0856e0141062 replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents: 25386
diff changeset
   258
    val tab' = fold (Symtab.update_list (op =)) fmts (mode_tab prtabs mode);
25386
82b62fe11d7a remove_prtabs: tuned, avoid excessive garbage;
wenzelm
parents: 24612
diff changeset
   259
  in AList.update (op =) (mode, tab') prtabs end;
15753
eb014dfc57ee tuned extend_prtabs;
wenzelm
parents: 15574
diff changeset
   260
25386
82b62fe11d7a remove_prtabs: tuned, avoid excessive garbage;
wenzelm
parents: 24612
diff changeset
   261
fun remove_prtabs mode xprods prtabs =
82b62fe11d7a remove_prtabs: tuned, avoid excessive garbage;
wenzelm
parents: 24612
diff changeset
   262
  let
82b62fe11d7a remove_prtabs: tuned, avoid excessive garbage;
wenzelm
parents: 24612
diff changeset
   263
    val tab = mode_tab prtabs mode;
82b62fe11d7a remove_prtabs: tuned, avoid excessive garbage;
wenzelm
parents: 24612
diff changeset
   264
    val fmts = map_filter (fn xprod as SynExt.XProd (_, _, c, _) =>
82b62fe11d7a remove_prtabs: tuned, avoid excessive garbage;
wenzelm
parents: 24612
diff changeset
   265
      if null (Symtab.lookup_list tab c) then NONE
82b62fe11d7a remove_prtabs: tuned, avoid excessive garbage;
wenzelm
parents: 24612
diff changeset
   266
      else xprod_to_fmt xprod) xprods;
25393
0856e0141062 replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents: 25386
diff changeset
   267
    val tab' = fold (Symtab.remove_list (op =)) fmts tab;
0856e0141062 replaced extend_prtabs by update_prtabs (absorb duplicates);
wenzelm
parents: 25386
diff changeset
   268
  in AList.update (op =) (mode, tab') prtabs end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   269
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   270
fun merge_prtabs prtabs1 prtabs2 =
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   271
  let
19046
bc5c6c9b114e removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 18977
diff changeset
   272
    val modes = distinct (op =) (map fst (prtabs1 @ prtabs2));
18957
8c3abd63bce3 TableFun: renamed xxx_multi to xxx_list;
wenzelm
parents: 18857
diff changeset
   273
    fun merge m = (m, Symtab.merge_list (op =) (mode_tab prtabs1 m, mode_tab prtabs2 m));
12292
c4090cc2aa15 Symtab.merge_multi';
wenzelm
parents: 12252
diff changeset
   274
  in map merge modes end;
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   275
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   277
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   278
(** pretty term or typ asts **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   279
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   280
fun is_chain [Block (_, pr)] = is_chain pr
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   281
  | is_chain [Arg _] = true
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   282
  | is_chain _  = false;
506
e0ca460d6e51 improved show_brackets again - Trueprop does not create () any more.
nipkow
parents: 505
diff changeset
   283
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   284
fun pretty extern ctxt tabs trf tokentrf type_mode curried ast0 p0 =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
  let
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   286
    val {extern_class, extern_type, extern_const} = extern;
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   287
26707
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   288
    fun token_trans a x =
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   289
      (case tokentrf a of
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   290
        NONE =>
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   291
          if member (op =) SynExt.standard_token_classes a
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   292
          then SOME (Pretty.str x) else NONE
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   293
      | SOME f => SOME (f ctxt x));
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   294
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   295
    (*default applications: prefix / postfix*)
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   296
    val appT =
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   297
      if type_mode then TypeExt.tappl_ast_tr'
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   298
      else if curried then SynTrans.applC_ast_tr'
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   299
      else SynTrans.appl_ast_tr';
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   300
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   301
    fun synT _ ([], args) = ([], args)
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   302
      | synT markup (Arg p :: symbs, t :: args) =
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   303
          let val (Ts, args') = synT markup (symbs, args);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   304
          in (astT (t, p) @ Ts, args') end
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   305
      | synT markup (TypArg p :: symbs, t :: args) =
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   306
          let
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   307
            val (Ts, args') = synT markup (symbs, args);
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   308
          in
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   309
            if type_mode then (astT (t, p) @ Ts, args')
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   310
            else (pretty extern ctxt tabs trf tokentrf true curried t p @ Ts, args')
237
a7d3e712767a MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
wenzelm
parents: 62
diff changeset
   311
          end
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   312
      | synT markup (String s :: symbs, args) =
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   313
          let val (Ts, args') = synT markup (symbs, args);
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   314
          in (markup (Pretty.str s) :: Ts, args') end
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   315
      | synT markup (Space s :: symbs, args) =
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   316
          let val (Ts, args') = synT markup (symbs, args);
8457
c5eb14ba754c use Pretty.str / Pretty.raw_str;
wenzelm
parents: 6767
diff changeset
   317
          in (Pretty.str s :: Ts, args') end
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   318
      | synT markup (Block (i, bsymbs) :: symbs, args) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   319
          let
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   320
            val (bTs, args') = synT markup (bsymbs, args);
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   321
            val (Ts, args'') = synT markup (symbs, args');
14837
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
   322
            val T =
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
   323
              if i < 0 then Pretty.unbreakable (Pretty.block bTs)
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
   324
              else Pretty.blk (i, bTs);
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
   325
          in (T :: Ts, args'') end
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   326
      | synT markup (Break i :: symbs, args) =
14837
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
   327
          let
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   328
            val (Ts, args') = synT markup (symbs, args);
14837
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
   329
            val T = if i < 0 then Pretty.fbrk else Pretty.brk i;
827c68f8267c added pp_show_brackets; support unbreakable blocks;
wenzelm
parents: 14783
diff changeset
   330
          in (T :: Ts, args') end
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   331
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   332
    and parT markup (pr, args, p, p': int) = #1 (synT markup
554
c7d9018cc9e6 various minor internal changes;
wenzelm
parents: 506
diff changeset
   333
          (if p > p' orelse
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   334
            (! show_brackets andalso p' <> SynExt.max_pri andalso not (is_chain pr))
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   335
            then [Block (1, Space "(" :: pr @ [Space ")"])]
554
c7d9018cc9e6 various minor internal changes;
wenzelm
parents: 506
diff changeset
   336
            else pr, args))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   337
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   338
    and atomT a = a |> Lexicon.unmark
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   339
     {case_class = fn c => Pretty.mark (Markup.tclass c) (Pretty.str (extern_class c)),
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   340
      case_type = fn c => Pretty.mark (Markup.tycon c) (Pretty.str (extern_type c)),
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   341
      case_const = fn c => Pretty.mark (Markup.const c) (Pretty.str (extern_const c)),
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   342
      case_fixed = fn x => the (token_trans "_free" x),
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   343
      case_default = Pretty.str}
19374
ae4a225e0c1f pretty: late externing of consts (support authentic syntax);
wenzelm
parents: 19046
diff changeset
   344
ae4a225e0c1f pretty: late externing of consts (support authentic syntax);
wenzelm
parents: 19046
diff changeset
   345
    and prefixT (_, a, [], _) = [atomT a]
16611
edb368e2878f proper treatment of advanced trfuns: pass thy argument;
wenzelm
parents: 15973
diff changeset
   346
      | prefixT (c, _, args, p) = astT (appT (c, args), p)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   347
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   348
    and splitT 0 ([x], ys) = (x, ys)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   349
      | splitT 0 (rev_xs, ys) = (Ast.Appl (rev rev_xs), ys)
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   350
      | splitT n (rev_xs, y :: ys) = splitT (n - 1) (y :: rev_xs, ys)
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   351
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   352
    and combT (tup as (c, a, args, p)) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   353
      let
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   354
        val nargs = length args;
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   355
        val markup = a |> Lexicon.unmark
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   356
         {case_class = Pretty.mark o Markup.tclass,
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   357
          case_type = Pretty.mark o Markup.tycon,
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   358
          case_const = Pretty.mark o Markup.const,
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   359
          case_fixed = Pretty.mark o Markup.fixed,
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   360
          case_default = K I};
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   361
2701
348ec44248df split ast_of_term(T);
wenzelm
parents: 2507
diff changeset
   362
        (*find matching table entry, or print as prefix / postfix*)
6280
218733fb6987 tuned pretty format lookup;
wenzelm
parents: 6273
diff changeset
   363
        fun prnt ([], []) = prefixT tup
18957
8c3abd63bce3 TableFun: renamed xxx_multi to xxx_list;
wenzelm
parents: 18857
diff changeset
   364
          | prnt ([], tb :: tbs) = prnt (Symtab.lookup_list tb a, tbs)
6280
218733fb6987 tuned pretty format lookup;
wenzelm
parents: 6273
diff changeset
   365
          | prnt ((pr, n, p') :: prnps, tbs) =
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   366
              if nargs = n then parT markup (pr, args, p, p')
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   367
              else if nargs > n andalso not type_mode then
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   368
                astT (appT (splitT n ([c], args)), p)
6280
218733fb6987 tuned pretty format lookup;
wenzelm
parents: 6273
diff changeset
   369
              else prnt (prnps, tbs);
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   370
      in
26707
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   371
        (case tokentrT a args of
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   372
          SOME prt => [prt]
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   373
        | NONE => astT (apply_trans ctxt (trf a) args, p) handle Match => prnt ([], tabs))
2200
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   374
      end
2538977e94fa added print_mode: string list ref (order of printer tables);
wenzelm
parents: 1509
diff changeset
   375
26707
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   376
    and tokentrT a [Ast.Variable x] = token_trans a x
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   377
      | tokentrT _ _ = NONE
ddf6bab64b96 token translations: context dependent, result Pretty.T;
wenzelm
parents: 25393
diff changeset
   378
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   379
    and astT (c as Ast.Constant a, p) = combT (c, a, [], p)
8457
c5eb14ba754c use Pretty.str / Pretty.raw_str;
wenzelm
parents: 6767
diff changeset
   380
      | astT (Ast.Variable x, _) = [Pretty.str x]
6280
218733fb6987 tuned pretty format lookup;
wenzelm
parents: 6273
diff changeset
   381
      | astT (Ast.Appl ((c as Ast.Constant a) :: (args as _ :: _)), p) = combT (c, a, args, p)
5691
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   382
      | astT (Ast.Appl (f :: (args as _ :: _)), p) = astT (appT (f, args), p)
3a6de95c09d0 no open;
wenzelm
parents: 4699
diff changeset
   383
      | astT (ast as Ast.Appl _, _) = raise Ast.AST ("pretty: malformed ast", [ast]);
6280
218733fb6987 tuned pretty format lookup;
wenzelm
parents: 6273
diff changeset
   384
  in astT (ast0, p0) end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   385
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   386
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   387
(* pretty_term_ast *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   388
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   389
fun pretty_term_ast extern ctxt curried prtabs trf tokentrf ast =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   390
  pretty extern ctxt (mode_tabs prtabs (print_mode_value ()))
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   391
    trf tokentrf false curried ast 0;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   392
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   393
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   394
(* pretty_typ_ast *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   395
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   396
fun pretty_typ_ast {extern_class, extern_type} ctxt _ prtabs trf tokentrf ast =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   397
  pretty {extern_class = extern_class, extern_type = extern_type, extern_const = I}
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   398
    ctxt (mode_tabs prtabs (print_mode_value ()))
23630
bc22daeed49e pretty: markup for syntax/name of authentic consts;
wenzelm
parents: 23615
diff changeset
   399
    trf tokentrf true false ast 0;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   400
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   401
end;