src/Pure/Syntax/type_ext.ML
author wenzelm
Mon, 31 May 2010 21:06:57 +0200
changeset 37216 3165bc303f66
parent 36502 586af36cf3cc
child 39288 f1ae2493d93f
permissions -rw-r--r--
modernized some structure names, keeping a few legacy aliases;
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
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
     4
Utilities for input and output of types.  The concrete syntax of types.
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 TYPE_EXT0 =
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
     8
sig
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
     9
  val sort_of_term: term -> sort
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    10
  val term_sorts: term -> (indexname * sort) list
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    11
  val typ_of_term: (indexname -> sort) -> term -> typ
27266
a2db1e379778 added type_constraint (formerly in type_infer.ML, which is now loaded later);
wenzelm
parents: 27197
diff changeset
    12
  val type_constraint: typ -> term -> term
22764
ccbd31bc1ef7 TypeExt.decode_term;
wenzelm
parents: 22704
diff changeset
    13
  val decode_term: (((string * int) * sort) list -> string * int -> sort) ->
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    14
    (string -> bool * string) -> (string -> string option) -> term -> term
6901
5e20ddfdf3c7 export term_of_typ;
wenzelm
parents: 5690
diff changeset
    15
  val term_of_typ: bool -> typ -> term
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
    16
  val no_brackets: unit -> bool
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    17
  val no_type_brackets: unit -> bool
35668
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
    18
  val type_ast_trs:
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
    19
   {read_class: Proof.context -> string -> string,
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
    20
    read_type: Proof.context -> string -> string} ->
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
    21
    (string * (Proof.context -> Ast.ast list -> Ast.ast)) list
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    22
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
signature TYPE_EXT =
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    25
sig
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
  include TYPE_EXT0
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    27
  val term_of_sort: sort -> term
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 764
diff changeset
    28
  val tappl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    29
  val sortT: typ
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36502
diff changeset
    30
  val type_ext: Syn_Ext.syn_ext
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    31
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36502
diff changeset
    33
structure Type_Ext: TYPE_EXT =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
struct
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    35
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    36
(** input utils **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    38
(* sort_of_term *)
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    39
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    40
fun sort_of_term tm =
557
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    41
  let
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    42
    fun err () = raise TERM ("sort_of_term: bad encoding of classes", [tm]);
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    43
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    44
    fun class s = Lexicon.unmark_class s handle Fail _ => err ();
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    45
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    46
    fun classes (Const (s, _)) = [class s]
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    47
      | classes (Const ("_classes", _) $ Const (s, _) $ cs) = class s :: classes cs
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    48
      | classes _ = err ();
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    50
    fun sort (Const ("_topsort", _)) = []
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    51
      | sort (Const (s, _)) = [class s]
3778
b70c41bc7491 fixed raw_term_sorts (again!);
wenzelm
parents: 2699
diff changeset
    52
      | sort (Const ("_sort", _) $ cs) = classes cs
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    53
      | sort _ = err ();
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    54
  in sort tm end;
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    55
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    56
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    57
(* term_sorts *)
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    58
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    59
fun term_sorts tm =
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    60
  let
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    61
    val sort_of = sort_of_term;
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    62
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    63
    fun add_env (Const ("_ofsort", _) $ 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", _) $ (Const ("_tfree", _) $ Free (x, _)) $ cs) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    66
          insert (op =) ((x, ~1), sort_of cs)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    67
      | add_env (Const ("_ofsort", _) $ 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 (Const ("_ofsort", _) $ (Const ("_tvar", _) $ Var (xi, _)) $ cs) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    70
          insert (op =) (xi, sort_of cs)
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    71
      | add_env (Abs (_, _, t)) = add_env t
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    72
      | add_env (t1 $ t2) = add_env t1 #> add_env t2
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
    73
      | add_env _ = I;
8895
2913a54e64cf added sort_of_term;
wenzelm
parents: 7500
diff changeset
    74
  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
    75
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    76
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
    77
(* typ_of_term *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    79
fun typ_of_term get_sort tm =
557
9d386e6c02b7 added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents: 347
diff changeset
    80
  let
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    81
    fun err () = raise TERM ("typ_of_term: bad encoding of type", [tm]);
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    82
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    83
    fun typ_of (Free (x, _)) = 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
    84
      | typ_of (Var (xi, _)) = TVar (xi, get_sort xi)
32784
1a5dde5079ac eliminated redundant bindings;
wenzelm
parents: 30364
diff changeset
    85
      | typ_of (Const ("_tfree",_) $ (t as Free _)) = typ_of t
1a5dde5079ac eliminated redundant bindings;
wenzelm
parents: 30364
diff changeset
    86
      | typ_of (Const ("_tvar",_) $ (t as Var _)) = typ_of t
12317
fed7bed97293 support "_::foo" sort constraints;
wenzelm
parents: 11312
diff changeset
    87
      | typ_of (Const ("_ofsort", _) $ Free (x, _) $ _) = TFree (x, get_sort (x, ~1))
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    88
      | typ_of (Const ("_ofsort", _) $ (Const ("_tfree",_) $ Free (x, _)) $ _) =
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    89
          TFree (x, get_sort (x, ~1))
12317
fed7bed97293 support "_::foo" sort constraints;
wenzelm
parents: 11312
diff changeset
    90
      | typ_of (Const ("_ofsort", _) $ Var (xi, _) $ _) = TVar (xi, get_sort xi)
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    91
      | typ_of (Const ("_ofsort", _) $ (Const ("_tvar",_) $ Var (xi, _)) $ _) =
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
    92
          TVar (xi, get_sort xi)
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    93
      | typ_of (Const ("_dummy_ofsort", _) $ t) = TFree ("'_dummy_", sort_of_term t)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    94
      | typ_of t =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
          let
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    96
            val (head, args) = Term.strip_comb t;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
            val a =
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    98
              (case head of
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
    99
                Const (c, _) => (Lexicon.unmark_type c handle Fail _ => err ())
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   100
              | _ => err ());
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   101
          in Type (a, map typ_of args) end;
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   102
  in typ_of tm end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
22764
ccbd31bc1ef7 TypeExt.decode_term;
wenzelm
parents: 22704
diff changeset
   105
(* decode_term -- transform parse tree into raw term *)
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   106
27266
a2db1e379778 added type_constraint (formerly in type_infer.ML, which is now loaded later);
wenzelm
parents: 27197
diff changeset
   107
fun type_constraint T t =
a2db1e379778 added type_constraint (formerly in type_infer.ML, which is now loaded later);
wenzelm
parents: 27197
diff changeset
   108
  if T = dummyT then t
a2db1e379778 added type_constraint (formerly in type_infer.ML, which is now loaded later);
wenzelm
parents: 27197
diff changeset
   109
  else Const ("_type_constraint_", T --> T) $ t;
a2db1e379778 added type_constraint (formerly in type_infer.ML, which is now loaded later);
wenzelm
parents: 27197
diff changeset
   110
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   111
fun decode_term get_sort map_const map_free tm =
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   112
  let
36448
edb757388592 get_sort: minimize sorts given in the text, while keeping those from the context unchanged (the latter are preferred);
wenzelm
parents: 35668
diff changeset
   113
    val decodeT = typ_of_term (get_sort (term_sorts tm));
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   114
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   115
    fun decode (Const ("_constrain", _) $ t $ typ) =
27266
a2db1e379778 added type_constraint (formerly in type_infer.ML, which is now loaded later);
wenzelm
parents: 27197
diff changeset
   116
          type_constraint (decodeT typ) (decode t)
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   117
      | decode (Const ("_constrainAbs", _) $ (Abs (x, T, t)) $ typ) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   118
          if T = dummyT then Abs (x, decodeT typ, decode t)
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   119
          else type_constraint (decodeT typ --> dummyT) (Abs (x, T, decode t))
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   120
      | decode (Abs (x, T, t)) = Abs (x, T, decode t)
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   121
      | decode (t $ u) = decode t $ decode u
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   122
      | decode (Const (a, T)) =
36502
586af36cf3cc uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
wenzelm
parents: 36448
diff changeset
   123
          (case try Lexicon.unmark_fixed a of
586af36cf3cc uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
wenzelm
parents: 36448
diff changeset
   124
            SOME x => Free (x, T)
586af36cf3cc uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
wenzelm
parents: 36448
diff changeset
   125
          | NONE =>
586af36cf3cc uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
wenzelm
parents: 36448
diff changeset
   126
              let val c =
586af36cf3cc uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
wenzelm
parents: 36448
diff changeset
   127
                (case try Lexicon.unmark_const a of
586af36cf3cc uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
wenzelm
parents: 36448
diff changeset
   128
                  SOME c => c
586af36cf3cc uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
wenzelm
parents: 36448
diff changeset
   129
                | NONE => snd (map_const a))
586af36cf3cc uniform decoding of fixed/const syntax entities, allows to pass "\<^fixed>foo__" through the syntax layer (supersedes 1b7109c10b7b);
wenzelm
parents: 36448
diff changeset
   130
              in Const (c, T) end)
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   131
      | decode (Free (a, T)) =
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   132
          (case (map_free a, map_const a) of
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   133
            (SOME x, _) => Free (x, T)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   134
          | (_, (true, c)) => Const (c, T)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   135
          | (_, (false, c)) => (if Long_Name.is_qualified c then Const else Free) (c, T))
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   136
      | decode (Var (xi, T)) = Var (xi, T)
22704
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   137
      | decode (t as Bound _) = t;
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   138
  in decode tm end;
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   139
f67607c3e56e added decode_types (from type_infer.ML);
wenzelm
parents: 20854
diff changeset
   140
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   141
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   142
(** output utils **)
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   143
2699
932fae4271d7 term_of_... now mark class, tfree, tvar;
wenzelm
parents: 2678
diff changeset
   144
(* term_of_sort *)
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   145
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   146
fun term_of_sort S =
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   147
  let
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   148
    val class = Lexicon.const o Lexicon.mark_class;
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   149
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   150
    fun classes [c] = class c
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   151
      | classes (c :: cs) = Lexicon.const "_classes" $ class c $ classes cs;
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   152
  in
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   153
    (case S of
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   154
      [] => Lexicon.const "_topsort"
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   155
    | [c] => class c
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   156
    | cs => Lexicon.const "_sort" $ classes cs)
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   157
  end;
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   158
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   159
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   160
(* term_of_typ *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   161
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   162
fun term_of_typ show_sorts ty =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
  let
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   164
    fun of_sort t S =
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   165
      if show_sorts then Lexicon.const "_ofsort" $ t $ term_of_sort S
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   166
      else t;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   167
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   168
    fun term_of (Type (a, Ts)) =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   169
          Term.list_comb (Lexicon.const (Lexicon.mark_type a), map term_of Ts)
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   170
      | term_of (TFree (x, S)) = of_sort (Lexicon.const "_tfree" $ Lexicon.free x) S
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   171
      | 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
   172
  in term_of ty end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   174
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
(** the type syntax **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   177
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   178
(* print mode *)
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   179
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   180
val bracketsN = "brackets";
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   181
val no_bracketsN = "no_brackets";
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   182
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   183
fun no_brackets () =
24613
bc889c3d55a3 added print_mode_value (CRITICAL);
wenzelm
parents: 23167
diff changeset
   184
  find_first (fn mode => mode = bracketsN orelse mode = no_bracketsN)
bc889c3d55a3 added print_mode_value (CRITICAL);
wenzelm
parents: 23167
diff changeset
   185
    (print_mode_value ()) = SOME no_bracketsN;
11312
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   186
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   187
val type_bracketsN = "type_brackets";
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   188
val no_type_bracketsN = "no_type_brackets";
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   189
4104bd8d1528 added (no)_type_brackets
nipkow
parents: 10572
diff changeset
   190
fun no_type_brackets () =
24613
bc889c3d55a3 added print_mode_value (CRITICAL);
wenzelm
parents: 23167
diff changeset
   191
  find_first (fn mode => mode = type_bracketsN orelse mode = no_type_bracketsN)
bc889c3d55a3 added print_mode_value (CRITICAL);
wenzelm
parents: 23167
diff changeset
   192
    (print_mode_value ()) <> SOME type_bracketsN;
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   193
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   194
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   195
(* parse ast translations *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   196
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   197
val class_ast = Ast.Constant o Lexicon.mark_class;
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   198
val type_ast = Ast.Constant o Lexicon.mark_type;
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   199
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   200
fun class_name_tr read_class (*"_class_name"*) [Ast.Variable c] = class_ast (read_class c)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   201
  | class_name_tr _ (*"_class_name"*) asts = raise Ast.AST ("class_name_tr", asts);
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   202
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   203
fun classes_tr read_class (*"_classes"*) [Ast.Variable c, ast] =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   204
      Ast.mk_appl (Ast.Constant "_classes") [class_ast (read_class c), ast]
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   205
  | classes_tr _ (*"_classes"*) asts = raise Ast.AST ("classes_tr", asts);
347
cd41a57221d0 changed translation of type applications according to new grammar;
wenzelm
parents: 330
diff changeset
   206
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   207
fun type_name_tr read_type (*"_type_name"*) [Ast.Variable c] = type_ast (read_type c)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   208
  | type_name_tr _ (*"_type_name"*) asts = raise Ast.AST ("type_name_tr", asts);
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   209
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   210
fun tapp_ast_tr read_type (*"_tapp"*) [ty, Ast.Variable c] =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   211
      Ast.Appl [type_ast (read_type c), ty]
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   212
  | tapp_ast_tr _ (*"_tapp"*) asts = raise Ast.AST ("tapp_ast_tr", asts);
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   213
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   214
fun tappl_ast_tr read_type (*"_tappl"*) [ty, tys, Ast.Variable c] =
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   215
      Ast.Appl (type_ast (read_type c) :: ty :: Ast.unfold_ast "_types" tys)
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   216
  | tappl_ast_tr _ (*"_tappl"*) asts = raise Ast.AST ("tappl_ast_tr", asts);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   217
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
fun bracket_ast_tr (*"_bracket"*) [dom, cod] =
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   219
      Ast.fold_ast_p "\\<^type>fun" (Ast.unfold_ast "_types" dom, cod)
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   220
  | bracket_ast_tr (*"_bracket"*) asts = raise Ast.AST ("bracket_ast_tr", asts);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   221
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   222
18
c9ec452ff08f lots of internal cleaning and tuning;
wenzelm
parents: 0
diff changeset
   223
(* print ast translations *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   224
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   225
fun tappl_ast_tr' (f, []) = raise Ast.AST ("tappl_ast_tr'", [f])
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   226
  | 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
   227
  | tappl_ast_tr' (f, ty :: tys) =
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   228
      Ast.Appl [Ast.Constant "_tappl", ty, Ast.fold_ast "_types" tys, f];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   229
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   230
fun fun_ast_tr' (*"\\<^type>fun"*) asts =
16614
a493a50e6c0a export no_type_brackets;
wenzelm
parents: 15833
diff changeset
   231
  if no_brackets () orelse no_type_brackets () then raise Match
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   232
  else
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   233
    (case Ast.unfold_ast_p "\\<^type>fun" (Ast.Appl (Ast.Constant "\\<^type>fun" :: asts)) of
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   234
      (dom as _ :: _ :: _, cod)
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   235
        => Ast.Appl [Ast.Constant "_bracket", Ast.fold_ast "_types" dom, cod]
b070825579b8 no_brackets mode;
wenzelm
parents: 9067
diff changeset
   236
    | _ => raise Match);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   237
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   238
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
(* type_ext *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   240
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   241
val sortT = Type ("sort", []);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   242
val classesT = Type ("classes", []);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   243
val typesT = Type ("types", []);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   244
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36502
diff changeset
   245
local open Lexicon Syn_Ext in
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   246
35668
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   247
val type_ext = syn_ext' false (K false)
347
cd41a57221d0 changed translation of type applications according to new grammar;
wenzelm
parents: 330
diff changeset
   248
  [Mfix ("_",           tidT --> typeT,                "", [], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   249
   Mfix ("_",           tvarT --> typeT,               "", [], max_pri),
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   250
   Mfix ("_",           idT --> typeT,                 "_type_name", [], max_pri),
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   251
   Mfix ("_",           longidT --> typeT,             "_type_name", [], max_pri),
330
2fda15dd1e0f changed the way a grammar is generated to allow the new parser to work;
clasohm
parents: 258
diff changeset
   252
   Mfix ("_::_",        [tidT, sortT] ---> typeT,      "_ofsort", [max_pri, 0], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   253
   Mfix ("_::_",        [tvarT, sortT] ---> typeT,     "_ofsort", [max_pri, 0], max_pri),
14838
b12855d44c97 tuned _dummy_ofsort syntax;
wenzelm
parents: 14255
diff changeset
   254
   Mfix ("'_()::_",     sortT --> typeT,               "_dummy_ofsort", [0], max_pri),
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   255
   Mfix ("_",           idT --> sortT,                 "_class_name", [], max_pri),
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   256
   Mfix ("_",           longidT --> sortT,             "_class_name", [], max_pri),
2584
b386951e15e6 improved comments;
wenzelm
parents: 2438
diff changeset
   257
   Mfix ("{}",          sortT,                         "_topsort", [], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   258
   Mfix ("{_}",         classesT --> sortT,            "_sort", [], max_pri),
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   259
   Mfix ("_",           idT --> classesT,              "_class_name", [], max_pri),
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   260
   Mfix ("_",           longidT --> classesT,          "_class_name", [], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   261
   Mfix ("_,_",         [idT, classesT] ---> classesT, "_classes", [], max_pri),
3829
d7333ef9e72c added longid syntax;
wenzelm
parents: 3778
diff changeset
   262
   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
   263
   Mfix ("_ _",         [typeT, idT] ---> typeT,       "_tapp", [max_pri, 0], max_pri),
3829
d7333ef9e72c added longid syntax;
wenzelm
parents: 3778
diff changeset
   264
   Mfix ("_ _",         [typeT, longidT] ---> typeT,   "_tapp", [max_pri, 0], max_pri),
9067
64464b5f6867 tuned tappl syntax;
wenzelm
parents: 8895
diff changeset
   265
   Mfix ("((1'(_,/ _')) _)", [typeT, typesT, idT] ---> typeT, "_tappl", [], max_pri),
64464b5f6867 tuned tappl syntax;
wenzelm
parents: 8895
diff changeset
   266
   Mfix ("((1'(_,/ _')) _)", [typeT, typesT, longidT] ---> typeT, "_tappl", [], max_pri),
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   267
   Mfix ("_",           typeT --> typesT,              "", [], max_pri),
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   268
   Mfix ("_,/ _",       [typeT, typesT] ---> typesT,   "_types", [], max_pri),
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   269
   Mfix ("(_/ => _)",   [typeT, typeT] ---> typeT,     "\\<^type>fun", [1, 0], 0),
624
33b9b5da3e6f made major changes to grammar;
clasohm
parents: 557
diff changeset
   270
   Mfix ("([_]/ => _)", [typesT, typeT] ---> typeT,    "_bracket", [0, 0], 0),
2678
d5fe793293ac added "_" syntax for dummyT;
wenzelm
parents: 2584
diff changeset
   271
   Mfix ("'(_')",       typeT --> typeT,               "", [0], max_pri),
35429
afa8cf9e63d8 authentic syntax for classes and type constructors;
wenzelm
parents: 35262
diff changeset
   272
   Mfix ("'_",          typeT,                         "\\<^type>dummy", [], max_pri)]
35433
73cc288b4f83 "_type_prop" is syntax const -- recovered printing of OFCLASS;
wenzelm
parents: 35429
diff changeset
   273
  ["_type_prop"]
37216
3165bc303f66 modernized some structure names, keeping a few legacy aliases;
wenzelm
parents: 36502
diff changeset
   274
  ([], [], [], map Syn_Ext.mk_trfun [("\\<^type>fun", K fun_ast_tr')])
26708
fc2e7d2f763d moved default token translations to proof_context.ML, if all fails the pretty printer falls back on plain output;
wenzelm
parents: 26039
diff changeset
   275
  []
239
08b6e842ec16 minor internal changes;
wenzelm
parents: 18
diff changeset
   276
  ([], []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   277
35668
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   278
fun type_ast_trs {read_class, read_type} =
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   279
 [("_class_name", class_name_tr o read_class),
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   280
  ("_classes", classes_tr o read_class),
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   281
  ("_type_name", type_name_tr o read_type),
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   282
  ("_tapp", tapp_ast_tr o read_type),
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   283
  ("_tappl", tappl_ast_tr o read_type),
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   284
  ("_bracket", K bracket_ast_tr)];
69e1740fbfb1 simplified Syntax.basic_syntax (again);
wenzelm
parents: 35433
diff changeset
   285
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   286
end;
5690
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   287
4b056ee5435c no open;
wenzelm
parents: 3829
diff changeset
   288
end;