| author | hoelzl | 
| Thu, 23 Jul 2015 16:40:47 +0200 | |
| changeset 60772 | a0cfa9050fa8 | 
| parent 59196 | 73a6403637b3 | 
| child 61476 | 1884c40f1539 | 
| permissions | -rw-r--r-- | 
| 18 | 1 | (* Title: Pure/Syntax/lexicon.ML | 
| 2 | Author: Tobias Nipkow and Markus Wenzel, TU Muenchen | |
| 0 | 3 | |
| 4703 | 4 | Lexer for the inner Isabelle syntax (terms and types). | 
| 18 | 5 | *) | 
| 0 | 6 | |
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 7 | signature LEXICON = | 
| 4247 
9bba9251bb4d
added implode_xstr: string list -> string, explode_xstr: string -> string list;
 wenzelm parents: 
3828diff
changeset | 8 | sig | 
| 42476 | 9 | structure Syntax: | 
| 10 | sig | |
| 11 | val const: string -> term | |
| 12 | val free: string -> term | |
| 13 | val var: indexname -> term | |
| 14 | end | |
| 30573 | 15 | val scan_id: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | 
| 16 | val scan_longid: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | |
| 17 | val scan_tid: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | |
| 18 | val scan_nat: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | |
| 40290 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 wenzelm parents: 
39510diff
changeset | 19 | val scan_float: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | 
| 30573 | 20 | val scan_hex: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | 
| 21 | val scan_bin: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | |
| 22 | val scan_var: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | |
| 23 | val scan_tvar: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list | |
| 50239 | 24 | val is_tid: string -> bool | 
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 25 | datatype token_kind = | 
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 26 | Literal | IdentSy | LongIdentSy | VarSy | TFreeSy | TVarSy | NumSy | | 
| 58421 | 27 | FloatSy | StrSy | StringSy | Cartouche | Space | Comment | EOF | 
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 28 | datatype token = Token of token_kind * string * Position.range | 
| 27806 | 29 | val str_of_token: token -> string | 
| 30 | val pos_of_token: token -> Position.T | |
| 27887 | 31 | val is_proper: token -> bool | 
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 32 | val mk_eof: Position.T -> token | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 33 | val eof: token | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 34 | val is_eof: token -> bool | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 35 | val stopper: token Scan.stopper | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
164diff
changeset | 36 | val terminals: string list | 
| 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
164diff
changeset | 37 | val is_terminal: string -> bool | 
| 49821 
d15fe10593ff
clarified output token markup (see also bc22daeed49e);
 wenzelm parents: 
48992diff
changeset | 38 | val literal_markup: string -> Markup.T | 
| 44736 | 39 | val report_of_token: token -> Position.report | 
| 39510 
d9f5f01faa1b
Syntax.read_asts error: report token ranges within message -- no side-effect here;
 wenzelm parents: 
39507diff
changeset | 40 | val reported_token_range: Proof.context -> token -> string | 
| 18 | 41 | val matching_tokens: token * token -> bool | 
| 42 | val valued_token: token -> bool | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
164diff
changeset | 43 | val predef_term: string -> token option | 
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 44 | val implode_string: Symbol.symbol list -> string | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 45 | val explode_string: string * Position.T -> Symbol_Pos.T list | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 46 | val implode_str: Symbol.symbol list -> string | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 47 | val explode_str: string * Position.T -> Symbol_Pos.T list | 
| 30573 | 48 | val tokenize: Scan.lexicon -> bool -> Symbol_Pos.T list -> token list | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 49 | val read_indexname: string -> indexname | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 50 | val read_var: string -> term | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 51 | val read_variable: string -> indexname option | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 52 | val read_nat: string -> int option | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 53 | val read_int: string -> int option | 
| 58421 | 54 |   val read_num: string -> {radix: int, leading_zeros: int, value: int}
 | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 55 |   val read_float: string -> {mant: int, exp: int}
 | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 56 | val mark_class: string -> string val unmark_class: string -> string | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 57 | val mark_type: string -> string val unmark_type: string -> string | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 58 | val mark_const: string -> string val unmark_const: string -> string | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 59 | val mark_fixed: string -> string val unmark_fixed: string -> string | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 60 | val unmark: | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 61 |    {case_class: string -> 'a,
 | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 62 | case_type: string -> 'a, | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 63 | case_const: string -> 'a, | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 64 | case_fixed: string -> 'a, | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 65 | case_default: string -> 'a} -> string -> 'a | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 66 | val is_marked: string -> bool | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 67 | val dummy_type: term | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 68 | val fun_type: term | 
| 59196 | 69 | val pp_lexicon: Scan.lexicon -> Pretty.T | 
| 4247 
9bba9251bb4d
added implode_xstr: string list -> string, explode_xstr: string -> string list;
 wenzelm parents: 
3828diff
changeset | 70 | end; | 
| 0 | 71 | |
| 14679 | 72 | structure Lexicon: LEXICON = | 
| 0 | 73 | struct | 
| 74 | ||
| 42476 | 75 | (** syntaxtic terms **) | 
| 76 | ||
| 77 | structure Syntax = | |
| 78 | struct | |
| 79 | ||
| 80 | fun const c = Const (c, dummyT); | |
| 81 | fun free x = Free (x, dummyT); | |
| 82 | fun var xi = Var (xi, dummyT); | |
| 83 | ||
| 84 | end; | |
| 85 | ||
| 86 | ||
| 87 | ||
| 4703 | 88 | (** basic scanners **) | 
| 89 | ||
| 30573 | 90 | open Basic_Symbol_Pos; | 
| 27773 | 91 | |
| 55105 | 92 | val err_prefix = "Inner lexical error: "; | 
| 93 | ||
| 94 | fun !!! msg = Symbol_Pos.!!! (fn () => err_prefix ^ msg); | |
| 27773 | 95 | |
| 50239 | 96 | val scan_id = Symbol_Pos.scan_ident; | 
| 27773 | 97 | val scan_longid = scan_id @@@ (Scan.repeat1 ($$$ "." @@@ scan_id) >> flat); | 
| 98 | val scan_tid = $$$ "'" @@@ scan_id; | |
| 4703 | 99 | |
| 40525 
14a2e686bdac
eliminated slightly odd pervasive Symbol_Pos.symbol;
 wenzelm parents: 
40290diff
changeset | 100 | val scan_nat = Scan.many1 (Symbol.is_digit o Symbol_Pos.symbol); | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
55962diff
changeset | 101 | val scan_float = scan_nat @@@ $$$ "." @@@ scan_nat; | 
| 40525 
14a2e686bdac
eliminated slightly odd pervasive Symbol_Pos.symbol;
 wenzelm parents: 
40290diff
changeset | 102 | val scan_hex = $$$ "0" @@@ $$$ "x" @@@ Scan.many1 (Symbol.is_ascii_hex o Symbol_Pos.symbol); | 
| 27773 | 103 | val scan_bin = $$$ "0" @@@ $$$ "b" @@@ Scan.many1 (fn (s, _) => s = "0" orelse s = "1"); | 
| 58421 | 104 | val scan_num = scan_hex || scan_bin || scan_nat; | 
| 4703 | 105 | |
| 27773 | 106 | val scan_id_nat = scan_id @@@ Scan.optional ($$$ "." @@@ scan_nat) []; | 
| 107 | val scan_var = $$$ "?" @@@ scan_id_nat; | |
| 108 | val scan_tvar = $$$ "?" @@@ $$$ "'" @@@ scan_id_nat; | |
| 4703 | 109 | |
| 50239 | 110 | fun is_tid s = | 
| 111 | (case try (unprefix "'") s of | |
| 112 | SOME s' => Symbol_Pos.is_identifier s' | |
| 113 | | NONE => false); | |
| 114 | ||
| 4703 | 115 | |
| 116 | ||
| 18 | 117 | (** datatype token **) | 
| 0 | 118 | |
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 119 | datatype token_kind = | 
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 120 | Literal | IdentSy | LongIdentSy | VarSy | TFreeSy | TVarSy | NumSy | | 
| 58421 | 121 | FloatSy | StrSy | StringSy | Cartouche | Space | Comment | EOF; | 
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 122 | |
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 123 | datatype token = Token of token_kind * string * Position.range; | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 124 | |
| 27806 | 125 | fun str_of_token (Token (_, s, _)) = s; | 
| 126 | fun pos_of_token (Token (_, _, (pos, _))) = pos; | |
| 127 | ||
| 27887 | 128 | fun is_proper (Token (Space, _, _)) = false | 
| 129 | | is_proper (Token (Comment, _, _)) = false | |
| 130 | | is_proper _ = true; | |
| 131 | ||
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 132 | |
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 133 | (* stopper *) | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 134 | |
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 135 | fun mk_eof pos = Token (EOF, "", (pos, Position.none)); | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 136 | val eof = mk_eof Position.none; | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 137 | |
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 138 | fun is_eof (Token (EOF, _, _)) = true | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 139 | | is_eof _ = false; | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 140 | |
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 141 | val stopper = Scan.stopper (K eof) is_eof; | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 142 | |
| 0 | 143 | |
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
164diff
changeset | 144 | (* terminal arguments *) | 
| 0 | 145 | |
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 146 | val terminal_kinds = | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 147 |  [("id", IdentSy),
 | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 148 |   ("longid", LongIdentSy),
 | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 149 |   ("var", VarSy),
 | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 150 |   ("tid", TFreeSy),
 | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 151 |   ("tvar", TVarSy),
 | 
| 45703 
c7a13ce60161
renamed inner syntax categories "num" to "num_token" and "xnum" to "xnum_token";
 wenzelm parents: 
45666diff
changeset | 152 |   ("num_token", NumSy),
 | 
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 153 |   ("float_token", FloatSy),
 | 
| 55033 | 154 |   ("str_token", StrSy),
 | 
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 155 |   ("string_token", StringSy),
 | 
| 55033 | 156 |   ("cartouche", Cartouche)];
 | 
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 157 | |
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 158 | val terminals = map #1 terminal_kinds; | 
| 20664 | 159 | val is_terminal = member (op =) terminals; | 
| 237 
a7d3e712767a
MAJOR INTERNAL CHANGE: extend and merge operations of syntax tables
 wenzelm parents: 
164diff
changeset | 160 | |
| 0 | 161 | |
| 27887 | 162 | (* markup *) | 
| 163 | ||
| 49821 
d15fe10593ff
clarified output token markup (see also bc22daeed49e);
 wenzelm parents: 
48992diff
changeset | 164 | fun literal_markup s = | 
| 52189 | 165 | if Symbol.is_ascii_identifier s orelse exists Symbol.is_letter (Symbol.explode s) | 
| 166 | then Markup.literal | |
| 167 | else Markup.delimiter; | |
| 49821 
d15fe10593ff
clarified output token markup (see also bc22daeed49e);
 wenzelm parents: 
48992diff
changeset | 168 | |
| 27887 | 169 | val token_kind_markup = | 
| 51612 
6a1e40f9dd55
added var_position in analogy to longid_position, for typing reports on input;
 wenzelm parents: 
50242diff
changeset | 170 | fn TFreeSy => Markup.tfree | 
| 55033 | 171 | | TVarSy => Markup.tvar | 
| 172 | | NumSy => Markup.numeral | |
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49821diff
changeset | 173 | | FloatSy => Markup.numeral | 
| 55033 | 174 | | StrSy => Markup.inner_string | 
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 175 | | StringSy => Markup.inner_string | 
| 55033 | 176 | | Cartouche => Markup.inner_cartouche | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49821diff
changeset | 177 | | Comment => Markup.inner_comment | 
| 55033 | 178 | | _ => Markup.empty; | 
| 27887 | 179 | |
| 44736 | 180 | fun report_of_token (Token (kind, s, (pos, _))) = | 
| 49821 
d15fe10593ff
clarified output token markup (see also bc22daeed49e);
 wenzelm parents: 
48992diff
changeset | 181 | let val markup = if kind = Literal then literal_markup s else token_kind_markup kind | 
| 44736 | 182 | in (pos, markup) end; | 
| 27887 | 183 | |
| 39510 
d9f5f01faa1b
Syntax.read_asts error: report token ranges within message -- no side-effect here;
 wenzelm parents: 
39507diff
changeset | 184 | fun reported_token_range ctxt tok = | 
| 39168 
e3ac771235f7
report token range after inner parse error -- often provides important clues about misunderstanding concerning lexical phase;
 wenzelm parents: 
38474diff
changeset | 185 | if is_proper tok | 
| 50201 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 wenzelm parents: 
49821diff
changeset | 186 | then Context_Position.reported_text ctxt (pos_of_token tok) Markup.token_range "" | 
| 39510 
d9f5f01faa1b
Syntax.read_asts error: report token ranges within message -- no side-effect here;
 wenzelm parents: 
39507diff
changeset | 187 | else ""; | 
| 39168 
e3ac771235f7
report token range after inner parse error -- often provides important clues about misunderstanding concerning lexical phase;
 wenzelm parents: 
38474diff
changeset | 188 | |
| 27887 | 189 | |
| 18 | 190 | (* matching_tokens *) | 
| 0 | 191 | |
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 192 | fun matching_tokens (Token (Literal, x, _), Token (Literal, y, _)) = x = y | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 193 | | matching_tokens (Token (k, _, _), Token (k', _, _)) = k = k'; | 
| 0 | 194 | |
| 195 | ||
| 18 | 196 | (* valued_token *) | 
| 0 | 197 | |
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 198 | fun valued_token (Token (Literal, _, _)) = false | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 199 | | valued_token (Token (EOF, _, _)) = false | 
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 200 | | valued_token _ = true; | 
| 0 | 201 | |
| 202 | ||
| 18 | 203 | (* predef_term *) | 
| 0 | 204 | |
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 205 | fun predef_term s = | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 206 | (case AList.lookup (op =) terminal_kinds s of | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 207 | SOME sy => SOME (Token (sy, s, Position.no_range)) | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 208 | | NONE => NONE); | 
| 0 | 209 | |
| 210 | ||
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 211 | |
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 212 | (** string literals **) | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 213 | |
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 214 | fun explode_literal scan_body (str, pos) = | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 215 | (case Scan.read Symbol_Pos.stopper scan_body (Symbol_Pos.explode (str, pos)) of | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 216 | SOME ss => ss | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 217 | | _ => error (err_prefix ^ "malformed string literal " ^ quote str ^ Position.here pos)); | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 218 | |
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 219 | |
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 220 | (* string *) | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 221 | |
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 222 | val scan_string = Scan.trace (Symbol_Pos.scan_string_qq err_prefix) >> #2; | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 223 | val scan_string_body = Symbol_Pos.scan_string_qq err_prefix >> (#1 o #2); | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 224 | |
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 225 | fun implode_string ss = quote (implode (map (fn "\"" => "\\\"" | s => s) ss)); | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 226 | val explode_string = explode_literal scan_string_body; | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 227 | |
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 228 | |
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 229 | (* str *) | 
| 18 | 230 | |
| 14730 | 231 | val scan_chr = | 
| 55107 | 232 | $$ "\\" |-- $$$ "'" || | 
| 40525 
14a2e686bdac
eliminated slightly odd pervasive Symbol_Pos.symbol;
 wenzelm parents: 
40290diff
changeset | 233 | Scan.one | 
| 58854 | 234 | ((fn s => s <> "\\" andalso s <> "'" andalso Symbol.not_eof s) o | 
| 40525 
14a2e686bdac
eliminated slightly odd pervasive Symbol_Pos.symbol;
 wenzelm parents: 
40290diff
changeset | 235 | Symbol_Pos.symbol) >> single || | 
| 55107 | 236 | $$$ "'" --| Scan.ahead (~$$ "'"); | 
| 14730 | 237 | |
| 238 | val scan_str = | |
| 55106 | 239 | Scan.ahead ($$ "'" -- $$ "'") |-- | 
| 240 | !!! "unclosed string literal" | |
| 241 | ($$$ "'" @@@ $$$ "'" @@@ (Scan.repeat scan_chr >> flat) @@@ $$$ "'" @@@ $$$ "'"); | |
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 242 | |
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 243 | val scan_str_body = | 
| 55107 | 244 | Scan.ahead ($$ "'" |-- $$ "'") |-- | 
| 55106 | 245 | !!! "unclosed string literal" | 
| 246 | ($$ "'" |-- $$ "'" |-- (Scan.repeat scan_chr >> flat) --| $$ "'" --| $$ "'"); | |
| 14730 | 247 | |
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 248 | fun implode_str ss = enclose "''" "''" (implode (map (fn "'" => "\\'" | s => s) ss)); | 
| 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 249 | val explode_str = explode_literal scan_str_body; | 
| 18 | 250 | |
| 251 | ||
| 27806 | 252 | |
| 18 | 253 | (** tokenize **) | 
| 254 | ||
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 255 | fun token_leq (Token (_, s1, _), Token (_, s2, _)) = s1 <= s2; | 
| 30573 | 256 | fun token kind ss = Token (kind, Symbol_Pos.content ss, Symbol_Pos.range ss); | 
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 257 | |
| 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 258 | fun tokenize lex xids syms = | 
| 18 | 259 | let | 
| 260 | val scan_xid = | |
| 55962 
fbd0e768bc8f
special identifier "__" (i.e. empty name with internal suffix) serves as wild-card for completion;
 wenzelm parents: 
55624diff
changeset | 261 | (if xids then $$$ "_" @@@ scan_id || scan_id else scan_id) || | 
| 
fbd0e768bc8f
special identifier "__" (i.e. empty name with internal suffix) serves as wild-card for completion;
 wenzelm parents: 
55624diff
changeset | 262 | $$$ "_" @@@ $$$ "_"; | 
| 18 | 263 | |
| 550 
353eea6ec232
replaced id, var, tid, tvar by idT, varT, tidT, tvarT;
 wenzelm parents: 
376diff
changeset | 264 | val scan_val = | 
| 26007 
3760d3ff4cce
basic scanners: produce symbol list instead of imploded string;
 wenzelm parents: 
24630diff
changeset | 265 | scan_tvar >> token TVarSy || | 
| 
3760d3ff4cce
basic scanners: produce symbol list instead of imploded string;
 wenzelm parents: 
24630diff
changeset | 266 | scan_var >> token VarSy || | 
| 
3760d3ff4cce
basic scanners: produce symbol list instead of imploded string;
 wenzelm parents: 
24630diff
changeset | 267 | scan_tid >> token TFreeSy || | 
| 28904 | 268 | scan_float >> token FloatSy || | 
| 58421 | 269 | scan_num >> token NumSy || | 
| 26007 
3760d3ff4cce
basic scanners: produce symbol list instead of imploded string;
 wenzelm parents: 
24630diff
changeset | 270 | scan_longid >> token LongIdentSy || | 
| 
3760d3ff4cce
basic scanners: produce symbol list instead of imploded string;
 wenzelm parents: 
24630diff
changeset | 271 | scan_xid >> token IdentSy; | 
| 18 | 272 | |
| 27800 
df444ddeff56
datatype token: maintain range, tuned representation;
 wenzelm parents: 
27773diff
changeset | 273 | val scan_lit = Scan.literal lex >> token Literal; | 
| 550 
353eea6ec232
replaced id, var, tid, tvar by idT, varT, tidT, tvarT;
 wenzelm parents: 
376diff
changeset | 274 | |
| 4703 | 275 | val scan_token = | 
| 55105 | 276 | Symbol_Pos.scan_cartouche err_prefix >> token Cartouche || | 
| 277 | Symbol_Pos.scan_comment err_prefix >> token Comment || | |
| 27887 | 278 | Scan.max token_leq scan_lit scan_val || | 
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 279 | scan_string >> token StringSy || | 
| 27887 | 280 | scan_str >> token StrSy || | 
| 40525 
14a2e686bdac
eliminated slightly odd pervasive Symbol_Pos.symbol;
 wenzelm parents: 
40290diff
changeset | 281 | Scan.many1 (Symbol.is_blank o Symbol_Pos.symbol) >> token Space; | 
| 18 | 282 | in | 
| 27773 | 283 | (case Scan.error | 
| 30573 | 284 | (Scan.finite Symbol_Pos.stopper (Scan.repeat scan_token)) syms of | 
| 27887 | 285 | (toks, []) => toks | 
| 55108 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 wenzelm parents: 
55107diff
changeset | 286 | | (_, ss) => | 
| 55624 | 287 |         error ("Inner lexical error" ^ Position.here (#1 (Symbol_Pos.range ss)) ^
 | 
| 288 |           Markup.markup Markup.no_report ("\nat " ^ quote (Symbol_Pos.content ss))))
 | |
| 289 | ||
| 18 | 290 | end; | 
| 291 | ||
| 292 | ||
| 293 | ||
| 294 | (** scan variables **) | |
| 295 | ||
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 296 | (* scan_indexname *) | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 297 | |
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 298 | local | 
| 18 | 299 | |
| 27773 | 300 | val scan_vname = | 
| 18 | 301 | let | 
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 302 | fun nat n [] = n | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 303 | | nat n (c :: cs) = nat (n * 10 + (ord c - ord "0")) cs; | 
| 18 | 304 | |
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 305 | fun idxname cs ds = (implode (rev cs), nat 0 ds); | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 306 | fun chop_idx [] ds = idxname [] ds | 
| 50242 
56b9c792a98b
support for sub-structured identifier syntax (inactive);
 wenzelm parents: 
50239diff
changeset | 307 | | chop_idx (cs as (_ :: "\\<^sub>" :: _)) ds = idxname cs ds | 
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 308 | | chop_idx (c :: cs) ds = | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 309 | if Symbol.is_digit c then chop_idx cs (c :: ds) | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 310 | else idxname (c :: cs) ds; | 
| 18 | 311 | |
| 50239 | 312 | val scan = | 
| 313 | (scan_id >> map Symbol_Pos.symbol) -- | |
| 55107 | 314 | Scan.optional ($$ "." |-- scan_nat >> (nat 0 o map Symbol_Pos.symbol)) ~1; | 
| 18 | 315 | in | 
| 27773 | 316 | scan >> | 
| 317 | (fn (cs, ~1) => chop_idx (rev cs) [] | |
| 318 | | (cs, i) => (implode cs, i)) | |
| 18 | 319 | end; | 
| 320 | ||
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 321 | in | 
| 18 | 322 | |
| 55107 | 323 | val scan_indexname = $$ "'" |-- scan_vname >> (fn (x, i) => ("'" ^ x, i)) || scan_vname;
 | 
| 15443 
07f78cc82a73
indexname function now parses type variables as well; changed input
 berghofe parents: 
14981diff
changeset | 324 | |
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 325 | end; | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 326 | |
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 327 | |
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 328 | (* indexname *) | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 329 | |
| 20313 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 330 | fun read_indexname s = | 
| 30573 | 331 | (case Scan.read Symbol_Pos.stopper scan_indexname (Symbol_Pos.explode (s, Position.none)) of | 
| 15531 | 332 | SOME xi => xi | 
| 20313 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 333 |   | _ => error ("Lexical error in variable name: " ^ quote s));
 | 
| 18 | 334 | |
| 335 | ||
| 4703 | 336 | (* read_var *) | 
| 18 | 337 | |
| 4703 | 338 | fun read_var str = | 
| 18 | 339 | let | 
| 340 | val scan = | |
| 55107 | 341 | $$ "?" |-- scan_indexname --| Scan.ahead (Scan.one Symbol_Pos.is_eof) | 
| 42476 | 342 | >> Syntax.var || | 
| 58854 | 343 | Scan.many (Symbol.not_eof o Symbol_Pos.symbol) | 
| 42476 | 344 | >> (Syntax.free o implode o map Symbol_Pos.symbol); | 
| 30573 | 345 | in the (Scan.read Symbol_Pos.stopper scan (Symbol_Pos.explode (str, Position.none))) end; | 
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 346 | |
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 347 | |
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 348 | (* read_variable *) | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 349 | |
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 350 | fun read_variable str = | 
| 55107 | 351 | let val scan = $$ "?" |-- scan_indexname || scan_indexname | 
| 30573 | 352 | in Scan.read Symbol_Pos.stopper scan (Symbol_Pos.explode (str, Position.none)) end; | 
| 4587 | 353 | |
| 354 | ||
| 20313 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 355 | (* read numbers *) | 
| 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 356 | |
| 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 357 | local | 
| 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 358 | |
| 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 359 | fun nat cs = | 
| 40525 
14a2e686bdac
eliminated slightly odd pervasive Symbol_Pos.symbol;
 wenzelm parents: 
40290diff
changeset | 360 | Option.map (#1 o Library.read_int o map Symbol_Pos.symbol) | 
| 30573 | 361 | (Scan.read Symbol_Pos.stopper scan_nat cs); | 
| 5860 | 362 | |
| 20313 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 363 | in | 
| 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 364 | |
| 30573 | 365 | fun read_nat s = nat (Symbol_Pos.explode (s, Position.none)); | 
| 20313 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 366 | |
| 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 367 | fun read_int s = | 
| 30573 | 368 | (case Symbol_Pos.explode (s, Position.none) of | 
| 27773 | 369 |     ("-", _) :: cs => Option.map ~ (nat cs)
 | 
| 20313 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 370 | | cs => nat cs); | 
| 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 371 | |
| 
bf9101cc4385
renamed Syntax.indexname to Syntax.read_indexname;
 wenzelm parents: 
20165diff
changeset | 372 | end; | 
| 5860 | 373 | |
| 374 | ||
| 58421 | 375 | (* read_num: hex/bin/decimal *) | 
| 9326 | 376 | |
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 377 | local | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 378 | |
| 20096 | 379 | val ten = ord "0" + 10; | 
| 380 | val a = ord "a"; | |
| 381 | val A = ord "A"; | |
| 35428 | 382 | val _ = a > A orelse raise Fail "Bad ASCII"; | 
| 20096 | 383 | |
| 384 | fun remap_hex c = | |
| 385 | let val x = ord c in | |
| 386 | if x >= a then chr (x - a + ten) | |
| 387 | else if x >= A then chr (x - A + ten) | |
| 388 | else c | |
| 389 | end; | |
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 390 | |
| 21781 | 391 | fun leading_zeros ["0"] = 0 | 
| 392 |   | leading_zeros ("0" :: cs) = 1 + leading_zeros cs
 | |
| 393 | | leading_zeros _ = 0; | |
| 394 | ||
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 395 | in | 
| 15965 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 paulson parents: 
15570diff
changeset | 396 | |
| 58421 | 397 | fun read_num str = | 
| 9326 | 398 | let | 
| 58421 | 399 | val (radix, digs) = | 
| 400 | (case Symbol.explode str of | |
| 401 | "0" :: "x" :: cs => (16, map remap_hex cs) | |
| 402 | | "0" :: "b" :: cs => (2, cs) | |
| 403 | | cs => (10, cs)); | |
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 404 | in | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 405 |    {radix = radix,
 | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 406 | leading_zeros = leading_zeros digs, | 
| 58421 | 407 | value = #1 (Library.read_radix_int radix digs)} | 
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 408 | end; | 
| 9326 | 409 | |
| 15991 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 410 | end; | 
| 
670f8e4b5a98
added read_variable: optional question mark on input;
 wenzelm parents: 
15965diff
changeset | 411 | |
| 28904 | 412 | fun read_float str = | 
| 413 | let | |
| 58421 | 414 | val cs = Symbol.explode str; | 
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 415 | val (intpart, fracpart) = | 
| 28904 | 416 | (case take_prefix Symbol.is_digit cs of | 
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 417 | (intpart, "." :: fracpart) => (intpart, fracpart) | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 418 | | _ => raise Fail "read_float"); | 
| 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 419 | in | 
| 58421 | 420 |    {mant = #1 (Library.read_int (intpart @ fracpart)),
 | 
| 29156 
89f76a58a378
renamed terminal category "float" to "float_token", to avoid name
 wenzelm parents: 
28904diff
changeset | 421 | exp = length fracpart} | 
| 42046 
6341c23baf10
added Lexicon.encode_position, Lexicon.decode_position;
 wenzelm parents: 
40525diff
changeset | 422 | end; | 
| 
6341c23baf10
added Lexicon.encode_position, Lexicon.decode_position;
 wenzelm parents: 
40525diff
changeset | 423 | |
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 424 | |
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 425 | (* marked logical entities *) | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 426 | |
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 427 | fun marker s = (prefix s, unprefix s); | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 428 | |
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 429 | val (mark_class, unmark_class) = marker "\\<^class>"; | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 430 | val (mark_type, unmark_type) = marker "\\<^type>"; | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 431 | val (mark_const, unmark_const) = marker "\\<^const>"; | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 432 | val (mark_fixed, unmark_fixed) = marker "\\<^fixed>"; | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 433 | |
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 434 | fun unmark {case_class, case_type, case_const, case_fixed, case_default} s =
 | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 435 | (case try unmark_class s of | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 436 | SOME c => case_class c | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 437 | | NONE => | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 438 | (case try unmark_type s of | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 439 | SOME c => case_type c | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 440 | | NONE => | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 441 | (case try unmark_const s of | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 442 | SOME c => case_const c | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 443 | | NONE => | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 444 | (case try unmark_fixed s of | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 445 | SOME c => case_fixed c | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 446 | | NONE => case_default s)))); | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 447 | |
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 448 | val is_marked = | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 449 |   unmark {case_class = K true, case_type = K true, case_const = K true,
 | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 450 | case_fixed = K true, case_default = K false}; | 
| 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 451 | |
| 42476 | 452 | val dummy_type = Syntax.const (mark_type "dummy"); | 
| 453 | val fun_type = Syntax.const (mark_type "fun"); | |
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
42264diff
changeset | 454 | |
| 59196 | 455 | |
| 456 | (* toplevel pretty printing *) | |
| 457 | ||
| 458 | val pp_lexicon = | |
| 459 |   Pretty.str_list "{" "}" o map quote o Scan.dest_lexicon;
 | |
| 460 | ||
| 0 | 461 | end; |