author | wenzelm |
Wed, 14 Jun 2000 18:00:46 +0200 | |
changeset 9067 | 64464b5f6867 |
parent 8895 | 2913a54e64cf |
child 10572 | b070825579b8 |
permissions | -rw-r--r-- |
18 | 1 |
(* Title: Pure/Syntax/type_ext.ML |
0 | 2 |
ID: $Id$ |
3 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen |
|
4 |
||
8895 | 5 |
Utilities for input and output of types. Also the concrete syntax of |
6 |
types, which is required to bootstrap Pure. |
|
0 | 7 |
*) |
8 |
||
9 |
signature TYPE_EXT0 = |
|
2584 | 10 |
sig |
8895 | 11 |
val sort_of_term: term -> sort |
3778 | 12 |
val raw_term_sorts: term -> (indexname * sort) list |
2584 | 13 |
val typ_of_term: (indexname -> sort) -> term -> typ |
6901 | 14 |
val term_of_typ: bool -> typ -> term |
2584 | 15 |
end; |
0 | 16 |
|
17 |
signature TYPE_EXT = |
|
2584 | 18 |
sig |
0 | 19 |
include TYPE_EXT0 |
2584 | 20 |
val term_of_sort: sort -> term |
1511 | 21 |
val tappl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast |
8895 | 22 |
val sortT: typ |
1511 | 23 |
val type_ext: SynExt.syn_ext |
2584 | 24 |
end; |
0 | 25 |
|
1511 | 26 |
structure TypeExt : TYPE_EXT = |
0 | 27 |
struct |
2584 | 28 |
|
29 |
||
30 |
(** input utils **) |
|
0 | 31 |
|
8895 | 32 |
(* sort_of_term *) |
2584 | 33 |
|
8895 | 34 |
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
|
35 |
let |
18 | 36 |
fun classes (Const (c, _)) = [c] |
37 |
| classes (Free (c, _)) = [c] |
|
3778 | 38 |
| classes (Const ("_classes", _) $ Const (c, _) $ cs) = c :: classes cs |
39 |
| classes (Const ("_classes", _) $ Free (c, _) $ cs) = c :: classes cs |
|
8895 | 40 |
| classes tm = raise TERM ("sort_of_term: bad encoding of classes", [tm]); |
0 | 41 |
|
2584 | 42 |
fun sort (Const ("_topsort", _)) = [] |
43 |
| sort (Const (c, _)) = [c] |
|
44 |
| sort (Free (c, _)) = [c] |
|
3778 | 45 |
| sort (Const ("_sort", _) $ cs) = classes cs |
8895 | 46 |
| sort tm = raise TERM ("sort_of_term: bad encoding of sort", [tm]); |
47 |
in sort tm end; |
|
48 |
||
2584 | 49 |
|
8895 | 50 |
(* raw_term_sorts *) |
51 |
||
52 |
fun raw_term_sorts tm = |
|
53 |
let |
|
54 |
fun add_env (Const ("_ofsort", _) $ Free (x, _) $ cs) env = ((x, ~1), sort_of_term cs) ins env |
|
55 |
| add_env (Const ("_ofsort", _) $ Var (xi, _) $ cs) env = (xi, sort_of_term cs) ins env |
|
3778 | 56 |
| add_env (Abs (_, _, t)) env = add_env t env |
57 |
| add_env (t1 $ t2) env = add_env t1 (add_env t2 env) |
|
58 |
| add_env t env = env; |
|
8895 | 59 |
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
|
60 |
|
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
61 |
|
2584 | 62 |
(* typ_of_term *) |
0 | 63 |
|
2584 | 64 |
fun typ_of_term get_sort t = |
557
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
65 |
let |
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
66 |
fun typ_of (Free (x, _)) = |
5690 | 67 |
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
|
68 |
else Type (x, []) |
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
69 |
| typ_of (Var (xi, _)) = TVar (xi, get_sort xi) |
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
70 |
| typ_of (Const ("_ofsort", _) $ Free (x, _) $ _) = |
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
71 |
TFree (x, get_sort (x, ~1)) |
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
72 |
| typ_of (Const ("_ofsort", _) $ Var (xi, _) $ _) = |
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
73 |
TVar (xi, get_sort xi) |
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
74 |
| typ_of tm = |
0 | 75 |
let |
76 |
val (t, ts) = strip_comb tm; |
|
77 |
val a = |
|
18 | 78 |
(case t of |
0 | 79 |
Const (x, _) => x |
80 |
| Free (x, _) => x |
|
3778 | 81 |
| _ => raise TERM ("typ_of_term: bad encoding of type", [tm])); |
0 | 82 |
in |
557
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
83 |
Type (a, map typ_of ts) |
9d386e6c02b7
added raw_term_sorts and changed typ_of_term accordingly (part of fix
wenzelm
parents:
347
diff
changeset
|
84 |
end; |
8895 | 85 |
in typ_of t end; |
0 | 86 |
|
87 |
||
88 |
||
2584 | 89 |
(** output utils **) |
90 |
||
2699 | 91 |
(* term_of_sort *) |
2584 | 92 |
|
93 |
fun term_of_sort S = |
|
94 |
let |
|
5690 | 95 |
fun class c = Lexicon.const "_class" $ Lexicon.free c; |
2584 | 96 |
|
97 |
fun classes [] = sys_error "term_of_sort" |
|
98 |
| classes [c] = class c |
|
5690 | 99 |
| classes (c :: cs) = Lexicon.const "_classes" $ class c $ classes cs; |
2584 | 100 |
in |
101 |
(case S of |
|
5690 | 102 |
[] => Lexicon.const "_topsort" |
2584 | 103 |
| [c] => class c |
5690 | 104 |
| cs => Lexicon.const "_sort" $ classes cs) |
2584 | 105 |
end; |
106 |
||
107 |
||
108 |
(* term_of_typ *) |
|
0 | 109 |
|
110 |
fun term_of_typ show_sorts ty = |
|
111 |
let |
|
2584 | 112 |
fun of_sort t S = |
5690 | 113 |
if show_sorts then Lexicon.const "_ofsort" $ t $ term_of_sort S |
2584 | 114 |
else t; |
0 | 115 |
|
5690 | 116 |
fun term_of (Type (a, Ts)) = list_comb (Lexicon.const a, map term_of Ts) |
117 |
| term_of (TFree (x, S)) = of_sort (Lexicon.const "_tfree" $ Lexicon.free x) S |
|
118 |
| term_of (TVar (xi, S)) = of_sort (Lexicon.const "_tvar" $ Lexicon.var xi) S; |
|
8895 | 119 |
in term_of ty end; |
0 | 120 |
|
121 |
||
122 |
||
123 |
(** the type syntax **) |
|
124 |
||
18 | 125 |
(* parse ast translations *) |
0 | 126 |
|
5690 | 127 |
fun tapp_ast_tr (*"_tapp"*) [ty, f] = Ast.Appl [f, ty] |
128 |
| 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
|
129 |
|
cd41a57221d0
changed translation of type applications according to new grammar;
wenzelm
parents:
330
diff
changeset
|
130 |
fun tappl_ast_tr (*"_tappl"*) [ty, tys, f] = |
5690 | 131 |
Ast.Appl (f :: ty :: Ast.unfold_ast "_types" tys) |
132 |
| tappl_ast_tr (*"_tappl"*) asts = raise Ast.AST ("tappl_ast_tr", asts); |
|
0 | 133 |
|
134 |
fun bracket_ast_tr (*"_bracket"*) [dom, cod] = |
|
5690 | 135 |
Ast.fold_ast_p "fun" (Ast.unfold_ast "_types" dom, cod) |
136 |
| bracket_ast_tr (*"_bracket"*) asts = raise Ast.AST ("bracket_ast_tr", asts); |
|
0 | 137 |
|
138 |
||
18 | 139 |
(* print ast translations *) |
0 | 140 |
|
5690 | 141 |
fun tappl_ast_tr' (f, []) = raise Ast.AST ("tappl_ast_tr'", [f]) |
142 |
| 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
|
143 |
| tappl_ast_tr' (f, ty :: tys) = |
5690 | 144 |
Ast.Appl [Ast.Constant "_tappl", ty, Ast.fold_ast "_types" tys, f]; |
0 | 145 |
|
146 |
fun fun_ast_tr' (*"fun"*) asts = |
|
5690 | 147 |
(case Ast.unfold_ast_p "fun" (Ast.Appl (Ast.Constant "fun" :: asts)) of |
18 | 148 |
(dom as _ :: _ :: _, cod) |
5690 | 149 |
=> Ast.Appl [Ast.Constant "_bracket", Ast.fold_ast "_types" dom, cod] |
0 | 150 |
| _ => raise Match); |
151 |
||
152 |
||
153 |
(* type_ext *) |
|
154 |
||
155 |
val sortT = Type ("sort", []); |
|
156 |
val classesT = Type ("classes", []); |
|
157 |
val typesT = Type ("types", []); |
|
158 |
||
5690 | 159 |
local open Lexicon SynExt in |
160 |
||
7500 | 161 |
val type_ext = mk_syn_ext false ["dummy"] |
347
cd41a57221d0
changed translation of type applications according to new grammar;
wenzelm
parents:
330
diff
changeset
|
162 |
[Mfix ("_", tidT --> typeT, "", [], max_pri), |
239 | 163 |
Mfix ("_", tvarT --> typeT, "", [], max_pri), |
164 |
Mfix ("_", idT --> typeT, "", [], max_pri), |
|
3829 | 165 |
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
|
166 |
Mfix ("_::_", [tidT, sortT] ---> typeT, "_ofsort", [max_pri, 0], max_pri), |
239 | 167 |
Mfix ("_::_", [tvarT, sortT] ---> typeT, "_ofsort", [max_pri, 0], max_pri), |
168 |
Mfix ("_", idT --> sortT, "", [], max_pri), |
|
3829 | 169 |
Mfix ("_", longidT --> sortT, "", [], max_pri), |
2584 | 170 |
Mfix ("{}", sortT, "_topsort", [], max_pri), |
239 | 171 |
Mfix ("{_}", classesT --> sortT, "_sort", [], max_pri), |
172 |
Mfix ("_", idT --> classesT, "", [], max_pri), |
|
3829 | 173 |
Mfix ("_", longidT --> classesT, "", [], max_pri), |
239 | 174 |
Mfix ("_,_", [idT, classesT] ---> classesT, "_classes", [], max_pri), |
3829 | 175 |
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
|
176 |
Mfix ("_ _", [typeT, idT] ---> typeT, "_tapp", [max_pri, 0], max_pri), |
3829 | 177 |
Mfix ("_ _", [typeT, longidT] ---> typeT, "_tapp", [max_pri, 0], max_pri), |
9067 | 178 |
Mfix ("((1'(_,/ _')) _)", [typeT, typesT, idT] ---> typeT, "_tappl", [], max_pri), |
179 |
Mfix ("((1'(_,/ _')) _)", [typeT, typesT, longidT] ---> typeT, "_tappl", [], max_pri), |
|
239 | 180 |
Mfix ("_", typeT --> typesT, "", [], max_pri), |
181 |
Mfix ("_,/ _", [typeT, typesT] ---> typesT, "_types", [], max_pri), |
|
182 |
Mfix ("(_/ => _)", [typeT, typeT] ---> typeT, "fun", [1, 0], 0), |
|
624 | 183 |
Mfix ("([_]/ => _)", [typesT, typeT] ---> typeT, "_bracket", [0, 0], 0), |
2678 | 184 |
Mfix ("'(_')", typeT --> typeT, "", [0], max_pri), |
185 |
Mfix ("'_", typeT, "dummy", [], max_pri)] |
|
258 | 186 |
[] |
347
cd41a57221d0
changed translation of type applications according to new grammar;
wenzelm
parents:
330
diff
changeset
|
187 |
([("_tapp", tapp_ast_tr), ("_tappl", tappl_ast_tr), ("_bracket", bracket_ast_tr)], |
239 | 188 |
[], |
189 |
[], |
|
190 |
[("fun", fun_ast_tr')]) |
|
2699 | 191 |
TokenTrans.token_translation |
239 | 192 |
([], []); |
0 | 193 |
|
194 |
end; |
|
5690 | 195 |
|
196 |
end; |