src/Pure/Syntax/type_ext.ML
author wenzelm
Sat, 07 Jul 2007 12:16:19 +0200
changeset 23631 2a9e918653cc
parent 23167 b9bbdf7eab3b
child 24613 bc889c3d55a3
permissions -rw-r--r--
pretty_sort/typ/term: markup;
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/type_ext.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
     5
Utilities for input and output of types.  Also the concrete syntax of
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
     6
types, which is required to bootstrap Pure.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
signature TYPE_EXT0 =
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    10
sig
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    11
  val sort_of_term: (sort -> sort) -> term -> sort
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    12
  val term_sorts: (sort -> sort) -> term -> (indexname * sort) list
12317
fed7bed97293 support "_::foo" sort constraints;
wenzelm
parents: 11312
diff changeset
    13
  val typ_of_term: (indexname -> sort) -> (sort -> sort) -> term -> typ
22764
ccbd31bc1ef7 TypeExt.decode_term;
wenzelm
parents: 22704
diff changeset
    14
  val decode_term: (((string * int) * sort) list -> string * int -> sort) ->
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    15
    (string -> string option) -> (string -> string option) ->
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    16
    (typ -> typ) -> (sort -> sort) -> term -> term
6901
5e20ddfdf3c7 export term_of_typ;
wenzelm
parents: 5690
diff changeset
    17
  val term_of_typ: bool -> typ -> term
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
    18
  val no_brackets: unit -> bool
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    19
  val no_type_brackets: unit -> bool
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    20
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
signature TYPE_EXT =
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    23
sig
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
  include TYPE_EXT0
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    25
  val term_of_sort: sort -> term
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 764
diff changeset
    26
  val tappl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    27
  val sortT: typ
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 764
diff changeset
    28
  val type_ext: SynExt.syn_ext
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    29
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
15750
860c282e6ca6 Syntax.mk_trfun;
wenzelm
parents: 15531
diff changeset
    31
structure TypeExt: TYPE_EXT =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
struct
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    33
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    34
(** input utils **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    36
(* sort_of_term *)
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    37
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    38
fun sort_of_term (map_sort: sort -> sort) tm =
557
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    39
  let
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
    40
    fun classes (Const (c, _)) = [c]
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
    41
      | classes (Free (c, _)) = [c]
3778
b70c41bc7491 fixed raw_term_sorts (again!);
wenzelm
parents: 2699
diff changeset
    42
      | classes (Const ("_classes", _) $ Const (c, _) $ cs) = c :: classes cs
b70c41bc7491 fixed raw_term_sorts (again!);
wenzelm
parents: 2699
diff changeset
    43
      | classes (Const ("_classes", _) $ Free (c, _) $ cs) = c :: classes cs
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    44
      | classes tm = raise TERM ("sort_of_term: bad encoding of classes", [tm]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    46
    fun sort (Const ("_topsort", _)) = []
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    47
      | sort (Const (c, _)) = [c]
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    48
      | sort (Free (c, _)) = [c]
14255
e6e3e3f0deed Records:
schirmer
parents: 14086
diff changeset
    49
      | sort (Const ("_class",_) $ Free (c, _)) = [c]
3778
b70c41bc7491 fixed raw_term_sorts (again!);
wenzelm
parents: 2699
diff changeset
    50
      | sort (Const ("_sort", _) $ cs) = classes cs
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    51
      | sort tm = raise TERM ("sort_of_term: bad encoding of sort", [tm]);
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    52
  in map_sort (sort tm) end;
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    53
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    54
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    55
(* term_sorts *)
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    56
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    57
fun term_sorts map_sort tm =
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    58
  let
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    59
    val sort_of = sort_of_term map_sort;
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    60
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    61
    fun add_env (Const ("_ofsort", _) $ Free (x, _) $ cs) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    62
          insert (op =) ((x, ~1), sort_of cs)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    63
      | add_env (Const ("_ofsort", _) $ (Const ("_tfree", _) $ Free (x, _)) $ cs) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    64
          insert (op =) ((x, ~1), sort_of cs)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    65
      | add_env (Const ("_ofsort", _) $ Var (xi, _) $ cs) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    66
          insert (op =) (xi, sort_of cs)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    67
      | add_env (Const ("_ofsort", _) $ (Const ("_tvar", _) $ Var (xi, _)) $ cs) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    68
          insert (op =) (xi, sort_of cs)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    69
      | add_env (Abs (_, _, t)) = add_env t
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    70
      | add_env (t1 $ t2) = add_env t1 #> add_env t2
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    71
      | add_env _ = I;
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    72
  in add_env tm [] end;
557
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    73
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    74
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    75
(* typ_of_term *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
12317
fed7bed97293 support "_::foo" sort constraints;
wenzelm
parents: 11312
diff changeset
    77
fun typ_of_term get_sort map_sort t =
557
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    78
  let
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    79
    fun typ_of (Free (x, _)) =
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
    80
          if Lexicon.is_tid x then TFree (x, get_sort (x, ~1))
557
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    81
          else Type (x, [])
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    82
      | typ_of (Var (xi, _)) = TVar (xi, get_sort xi)
14255
e6e3e3f0deed Records:
schirmer
parents: 14086
diff changeset
    83
      | typ_of (Const ("_tfree",_) $ (t as Free x)) = typ_of t
e6e3e3f0deed Records:
schirmer
parents: 14086
diff changeset
    84
      | typ_of (Const ("_tvar",_) $ (t as Var x)) = typ_of t
12317
fed7bed97293 support "_::foo" sort constraints;
wenzelm
parents: 11312
diff changeset
    85
      | typ_of (Const ("_ofsort", _) $ Free (x, _) $ _) = TFree (x, get_sort (x, ~1))
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    86
      | typ_of (Const ("_ofsort", _) $ (Const ("_tfree",_) $ Free (x, _)) $ _) =
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    87
          TFree (x, get_sort (x, ~1))
12317
fed7bed97293 support "_::foo" sort constraints;
wenzelm
parents: 11312
diff changeset
    88
      | typ_of (Const ("_ofsort", _) $ Var (xi, _) $ _) = TVar (xi, get_sort xi)
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    89
      | typ_of (Const ("_ofsort", _) $ (Const ("_tvar",_) $ Var (xi, _)) $ _) =
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    90
          TVar (xi, get_sort xi)
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    91
      | typ_of (Const ("_dummy_ofsort", _) $ t) = TFree ("'_dummy_", sort_of_term map_sort t)
557
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    92
      | typ_of tm =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
          let
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    94
            val (t, ts) = Term.strip_comb tm;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
            val a =
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
    96
              (case t of
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
                Const (x, _) => x
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
              | Free (x, _) => x
3778
b70c41bc7491 fixed raw_term_sorts (again!);
wenzelm
parents: 2699
diff changeset
    99
              | _ => raise TERM ("typ_of_term: bad encoding of type", [tm]));
12317
fed7bed97293 support "_::foo" sort constraints;
wenzelm
parents: 11312
diff changeset
   100
          in Type (a, map typ_of ts) end;
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
   101
  in typ_of t end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
22764
ccbd31bc1ef7 TypeExt.decode_term;
wenzelm
parents: 22704
diff changeset
   104
(* decode_term -- transform parse tree into raw term *)
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   105
22764
ccbd31bc1ef7 TypeExt.decode_term;
wenzelm
parents: 22704
diff changeset
   106
fun decode_term get_sort map_const map_free map_type map_sort tm =
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   107
  let
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   108
    val sort_env = term_sorts map_sort tm;
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   109
    val decodeT = map_type o typ_of_term (get_sort sort_env) map_sort;
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   110
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   111
    fun decode (Const ("_constrain", _) $ t $ typ) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   112
          TypeInfer.constrain (decode t) (decodeT typ)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   113
      | decode (Const ("_constrainAbs", _) $ (Abs (x, T, t)) $ typ) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   114
          if T = dummyT then Abs (x, decodeT typ, decode t)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   115
          else TypeInfer.constrain (Abs (x, map_type T, decode t)) (decodeT typ --> dummyT)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   116
      | decode (Abs (x, T, t)) = Abs (x, map_type T, decode t)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   117
      | decode (t $ u) = decode t $ decode u
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   118
      | decode (Const (a, T)) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   119
          let val c =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   120
            (case try (unprefix Lexicon.constN) a of SOME c => c | NONE => perhaps map_const a)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   121
          in Const (c, map_type T) end
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   122
      | decode (Free (a, T)) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   123
          (case (map_free a, map_const a) of
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   124
            (SOME x, _) => Free (x, map_type T)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   125
          | (_, SOME c) => Const (c, map_type T)
23167
b9bbdf7eab3b decode_term: force qualified name into Const;
wenzelm
parents: 22764
diff changeset
   126
          | _ => (if NameSpace.is_qualified a then Const else Free) (a, map_type T))
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   127
      | decode (Var (xi, T)) = Var (xi, map_type T)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   128
      | decode (t as Bound _) = t;
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   129
  in decode tm end;
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   130
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   131
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   133
(** output utils **)
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   134
2699
932fae4271d7 term_of_... now mark class, tfree, tvar;
wenzelm
parents: 2678
diff changeset
   135
(* term_of_sort *)
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   136
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   137
fun term_of_sort S =
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   138
  let
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   139
    fun class c = Lexicon.const "_class" $ Lexicon.free c;
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   140
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   141
    fun classes [] = sys_error "term_of_sort"
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   142
      | classes [c] = class c
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   143
      | classes (c :: cs) = Lexicon.const "_classes" $ class c $ classes cs;
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   144
  in
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   145
    (case S of
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   146
      [] => Lexicon.const "_topsort"
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   147
    | [c] => class c
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   148
    | cs => Lexicon.const "_sort" $ classes cs)
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   149
  end;
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   150
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   151
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   152
(* term_of_typ *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   153
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
fun term_of_typ show_sorts ty =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   155
  let
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   156
    fun of_sort t S =
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   157
      if show_sorts then Lexicon.const "_ofsort" $ t $ term_of_sort S
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   158
      else t;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   159
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
   160
    fun term_of (Type (a, Ts)) = Term.list_comb (Lexicon.const a, map term_of Ts)
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   161
      | term_of (TFree (x, S)) = of_sort (Lexicon.const "_tfree" $ Lexicon.free x) S
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   162
      | term_of (TVar (xi, S)) = of_sort (Lexicon.const "_tvar" $ Lexicon.var xi) S;
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
   163
  in term_of ty end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   165
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   167
(** the type syntax **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   168
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   169
(* print mode *)
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   170
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   171
val bracketsN = "brackets";
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   172
val no_bracketsN = "no_brackets";
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   173
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   174
fun no_brackets () =
19305
5c16895d548b avoid polymorphic equality;
wenzelm
parents: 16614
diff changeset
   175
  find_first (fn mode => mode = bracketsN orelse mode = no_bracketsN) (! print_mode)
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
   176
    = SOME no_bracketsN;
11312
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   177
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   178
val type_bracketsN = "type_brackets";
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   179
val no_type_bracketsN = "no_type_brackets";
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   180
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   181
fun no_type_brackets () =
19305
5c16895d548b avoid polymorphic equality;
wenzelm
parents: 16614
diff changeset
   182
  find_first (fn mode => mode = type_bracketsN orelse mode = no_type_bracketsN) (! print_mode)
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
   183
    <> SOME type_bracketsN;
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   184
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   185
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   186
(* parse ast translations *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   188
fun tapp_ast_tr (*"_tapp"*) [ty, f] = Ast.Appl [f, ty]
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   189
  | tapp_ast_tr (*"_tapp"*) asts = raise Ast.AST ("tapp_ast_tr", asts);
347
cd41a57221d0 changed translation of type applications according to new grammar;
wenzelm
parents: 330
diff changeset
   190
cd41a57221d0 changed translation of type applications according to new grammar;
wenzelm
parents: 330
diff changeset
   191
fun tappl_ast_tr (*"_tappl"*) [ty, tys, f] =
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   192
      Ast.Appl (f :: ty :: Ast.unfold_ast "_types" tys)
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   193
  | tappl_ast_tr (*"_tappl"*) asts = raise Ast.AST ("tappl_ast_tr", asts);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   195
fun bracket_ast_tr (*"_bracket"*) [dom, cod] =
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   196
      Ast.fold_ast_p "fun" (Ast.unfold_ast "_types" dom, cod)
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   197
  | bracket_ast_tr (*"_bracket"*) asts = raise Ast.AST ("bracket_ast_tr", asts);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   198
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   200
(* print ast translations *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   201
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   202
fun tappl_ast_tr' (f, []) = raise Ast.AST ("tappl_ast_tr'", [f])
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   203
  | tappl_ast_tr' (f, [ty]) = Ast.Appl [Ast.Constant "_tapp", ty, f]
347
cd41a57221d0 changed translation of type applications according to new grammar;
wenzelm
parents: 330
diff changeset
   204
  | tappl_ast_tr' (f, ty :: tys) =
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   205
      Ast.Appl [Ast.Constant "_tappl", ty, Ast.fold_ast "_types" tys, f];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   206
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   207
fun fun_ast_tr' (*"fun"*) asts =
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
   208
  if no_brackets () orelse no_type_brackets () then raise Match
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   209
  else
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   210
    (case Ast.unfold_ast_p "fun" (Ast.Appl (Ast.Constant "fun" :: asts)) of
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   211
      (dom as _ :: _ :: _, cod)
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   212
        => Ast.Appl [Ast.Constant "_bracket", Ast.fold_ast "_types" dom, cod]
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   213
    | _ => raise Match);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   214
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   215
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
(* type_ext *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   217
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
val sortT = Type ("sort", []);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   219
val classesT = Type ("classes", []);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   220
val typesT = Type ("types", []);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   221
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   222
local open Lexicon SynExt in
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   223
14903
d264b8ad3eec removed separate logtypes field of syntax;
wenzelm
parents: 14838
diff changeset
   224
val type_ext = syn_ext' false (K false)
347
cd41a57221d0 changed translation of type applications according to new grammar;
wenzelm
parents: 330
diff changeset
   225
  [Mfix ("_",           tidT --> typeT,                "", [], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   226
   Mfix ("_",           tvarT --> typeT,               "", [], max_pri),
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   227
   Mfix ("_",           idT --> typeT,                 "", [], max_pri),
3829
d7333ef9e72c added longid syntax;
wenzelm
parents: 3778
diff changeset
   228
   Mfix ("_",           longidT --> typeT,             "", [], max_pri),
330
2fda15dd1e0f changed the way a grammar is generated to allow the new parser to work;
clasohm
parents: 258
diff changeset
   229
   Mfix ("_::_",        [tidT, sortT] ---> typeT,      "_ofsort", [max_pri, 0], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   230
   Mfix ("_::_",        [tvarT, sortT] ---> typeT,     "_ofsort", [max_pri, 0], max_pri),
14838
b12855d44c97 tuned _dummy_ofsort syntax;
wenzelm
parents: 14255
diff changeset
   231
   Mfix ("'_()::_",     sortT --> typeT,               "_dummy_ofsort", [0], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   232
   Mfix ("_",           idT --> sortT,                 "", [], max_pri),
3829
d7333ef9e72c added longid syntax;
wenzelm
parents: 3778
diff changeset
   233
   Mfix ("_",           longidT --> sortT,             "", [], max_pri),
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   234
   Mfix ("{}",          sortT,                         "_topsort", [], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   235
   Mfix ("{_}",         classesT --> sortT,            "_sort", [], max_pri),
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   236
   Mfix ("_",           idT --> classesT,              "", [], max_pri),
3829
d7333ef9e72c added longid syntax;
wenzelm
parents: 3778
diff changeset
   237
   Mfix ("_",           longidT --> classesT,          "", [], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   238
   Mfix ("_,_",         [idT, classesT] ---> classesT, "_classes", [], max_pri),
3829
d7333ef9e72c added longid syntax;
wenzelm
parents: 3778
diff changeset
   239
   Mfix ("_,_",         [longidT, classesT] ---> classesT, "_classes", [], max_pri),
330
2fda15dd1e0f changed the way a grammar is generated to allow the new parser to work;
clasohm
parents: 258
diff changeset
   240
   Mfix ("_ _",         [typeT, idT] ---> typeT,       "_tapp", [max_pri, 0], max_pri),
3829
d7333ef9e72c added longid syntax;
wenzelm
parents: 3778
diff changeset
   241
   Mfix ("_ _",         [typeT, longidT] ---> typeT,   "_tapp", [max_pri, 0], max_pri),
9067
64464b5f6867 tuned tappl syntax;
wenzelm
parents: 8895
diff changeset
   242
   Mfix ("((1'(_,/ _')) _)", [typeT, typesT, idT] ---> typeT, "_tappl", [], max_pri),
64464b5f6867 tuned tappl syntax;
wenzelm
parents: 8895
diff changeset
   243
   Mfix ("((1'(_,/ _')) _)", [typeT, typesT, longidT] ---> typeT, "_tappl", [], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   244
   Mfix ("_",           typeT --> typesT,              "", [], max_pri),
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   245
   Mfix ("_,/ _",       [typeT, typesT] ---> typesT,   "_types", [], max_pri),
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   246
   Mfix ("(_/ => _)",   [typeT, typeT] ---> typeT,     "fun", [1, 0], 0),
624
33b9b5da3e6f made major changes to grammar;
clasohm
parents: 557
diff changeset
   247
   Mfix ("([_]/ => _)", [typesT, typeT] ---> typeT,    "_bracket", [0, 0], 0),
2678
d5fe793293ac added "_" syntax for dummyT;
wenzelm
parents: 2584
diff changeset
   248
   Mfix ("'(_')",       typeT --> typeT,               "", [0], max_pri),
d5fe793293ac added "_" syntax for dummyT;
wenzelm
parents: 2584
diff changeset
   249
   Mfix ("'_",          typeT,                         "dummy", [], max_pri)]
258
e540b7d4ecb1 minor internal changes;
wenzelm
parents: 239
diff changeset
   250
  []
15750
860c282e6ca6 Syntax.mk_trfun;
wenzelm
parents: 15531
diff changeset
   251
  (map SynExt.mk_trfun
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
   252
   [("_tapp", K tapp_ast_tr), ("_tappl", K tappl_ast_tr), ("_bracket", K bracket_ast_tr)],
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   253
   [],
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   254
   [],
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
   255
   map SynExt.mk_trfun [("fun", K fun_ast_tr')])
15833
78109c7012ed removed token_trans.ML (some content moved to syn_ext.ML);
wenzelm
parents: 15750
diff changeset
   256
  (map (fn s => ("", s, Output.output_width)) SynExt.standard_token_classes)
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   257
  ([], []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   258
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
end;
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   260
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   261
end;