| author | wenzelm |
| Wed, 30 Mar 2016 22:35:41 +0200 | |
| changeset 62772 | 77bbe5af41c3 |
| parent 62764 | ff3b8e4079bd |
| child 62783 | 75ee05386b90 |
| permissions | -rw-r--r-- |
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42268
diff
changeset
|
1 |
(* Title: Pure/Syntax/syntax_ext.ML |
|
911
55754d6d399c
new in mixfix annotations: "' " (quote space) separates delimiters without
wenzelm
parents:
865
diff
changeset
|
2 |
Author: Markus Wenzel and Carsten Clasohm, TU Muenchen |
| 240 | 3 |
|
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42268
diff
changeset
|
4 |
Syntax extension. |
| 240 | 5 |
*) |
6 |
||
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42268
diff
changeset
|
7 |
signature SYNTAX_EXT = |
| 4050 | 8 |
sig |
| 62753 | 9 |
datatype mfix = Mfix of Symbol_Pos.T list * typ * string * int list * int * Position.T |
| 1510 | 10 |
val typ_to_nonterm: typ -> string |
11 |
datatype xsymb = |
|
12 |
Delim of string | |
|
13 |
Argument of string * int | |
|
14 |
Space of string | |
|
15 |
Bg of int | Brk of int | En |
|
16 |
datatype xprod = XProd of string * xsymb list * string * int |
|
17 |
val chain_pri: int |
|
| 4701 | 18 |
val delims_of: xprod list -> string list list |
| 1510 | 19 |
datatype syn_ext = |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35429
diff
changeset
|
20 |
Syn_Ext of {
|
| 1510 | 21 |
xprods: xprod list, |
|
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42297
diff
changeset
|
22 |
consts: (string * string) list, |
| 21772 | 23 |
parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list, |
| 1510 | 24 |
parse_rules: (Ast.ast * Ast.ast) list, |
| 21772 | 25 |
parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list, |
|
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
26 |
print_translation: (string * ((Proof.context -> typ -> term list -> term) * stamp)) list, |
| 1510 | 27 |
print_rules: (Ast.ast * Ast.ast) list, |
| 42268 | 28 |
print_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list} |
| 62752 | 29 |
val mfix_args: Symbol_Pos.T list -> int |
30 |
val mixfix_args: Input.source -> int |
|
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42268
diff
changeset
|
31 |
val escape: string -> string |
| 59841 | 32 |
val syn_ext': string list -> mfix list -> |
|
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42297
diff
changeset
|
33 |
(string * string) list -> (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * |
| 21772 | 34 |
(string * ((Proof.context -> term list -> term) * stamp)) list * |
|
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
35 |
(string * ((Proof.context -> typ -> term list -> term) * stamp)) list * |
| 42268 | 36 |
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> |
37 |
(Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext |
|
|
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42297
diff
changeset
|
38 |
val syn_ext: mfix list -> (string * string) list -> |
| 21772 | 39 |
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * |
40 |
(string * ((Proof.context -> term list -> term) * stamp)) list * |
|
|
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
41 |
(string * ((Proof.context -> typ -> term list -> term) * stamp)) list * |
| 42268 | 42 |
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> |
43 |
(Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext |
|
| 14903 | 44 |
val syn_ext_rules: (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext |
45 |
val syn_ext_trfuns: |
|
| 21772 | 46 |
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * |
47 |
(string * ((Proof.context -> term list -> term) * stamp)) list * |
|
|
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
48 |
(string * ((Proof.context -> typ -> term list -> term) * stamp)) list * |
| 21772 | 49 |
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> syn_ext |
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42268
diff
changeset
|
50 |
val stamp_trfun: stamp -> string * 'a -> string * ('a * stamp)
|
|
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42268
diff
changeset
|
51 |
val mk_trfun: string * 'a -> string * ('a * stamp)
|
|
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42268
diff
changeset
|
52 |
val eq_trfun: ('a * stamp) * ('a * stamp) -> bool
|
| 4050 | 53 |
end; |
| 240 | 54 |
|
|
42288
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
wenzelm
parents:
42268
diff
changeset
|
55 |
structure Syntax_Ext: SYNTAX_EXT = |
| 240 | 56 |
struct |
57 |
||
58 |
(** datatype xprod **) |
|
59 |
||
60 |
(*Delim s: delimiter s |
|
61 |
Argument (s, p): nonterminal s requiring priority >= p, or valued token |
|
62 |
Space s: some white space for printing |
|
63 |
Bg, Brk, En: blocks and breaks for pretty printing*) |
|
64 |
||
65 |
datatype xsymb = |
|
66 |
Delim of string | |
|
67 |
Argument of string * int | |
|
68 |
Space of string | |
|
69 |
Bg of int | Brk of int | En; |
|
70 |
||
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
71 |
fun is_delim (Delim _) = true |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
72 |
| is_delim _ = false; |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
73 |
|
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
74 |
fun is_terminal (Delim _) = true |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
75 |
| is_terminal (Argument (s, _)) = Lexicon.is_terminal s |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
76 |
| is_terminal _ = false; |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
77 |
|
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
78 |
fun is_argument (Argument _) = true |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
79 |
| is_argument _ = false; |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
80 |
|
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
81 |
fun is_index (Argument ("index", _)) = true
|
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
82 |
| is_index _ = false; |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
83 |
|
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
84 |
val index = Argument ("index", 1000);
|
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
85 |
|
| 240 | 86 |
|
87 |
(*XProd (lhs, syms, c, p): |
|
88 |
lhs: name of nonterminal on the lhs of the production |
|
89 |
syms: list of symbols on the rhs of the production |
|
90 |
c: head of parse tree |
|
91 |
p: priority of this production*) |
|
92 |
||
93 |
datatype xprod = XProd of string * xsymb list * string * int; |
|
94 |
||
95 |
val chain_pri = ~1; (*dummy for chain productions*) |
|
96 |
||
97 |
fun delims_of xprods = |
|
| 19004 | 98 |
fold (fn XProd (_, xsymbs, _, _) => |
99 |
fold (fn Delim s => insert (op =) s | _ => I) xsymbs) xprods [] |
|
100 |
|> map Symbol.explode; |
|
| 240 | 101 |
|
| 62764 | 102 |
fun reports_of (xsym, pos: Position.T) = |
103 |
(case xsym of |
|
104 |
Delim _ => [(pos, Markup.literal)] |
|
105 |
| Bg _ => [(pos, Markup.keyword3)] |
|
106 |
| Brk _ => [(pos, Markup.keyword3)] |
|
107 |
| En => [(pos, Markup.keyword3)] |
|
108 |
| _ => []); |
|
109 |
||
| 240 | 110 |
|
111 |
||
112 |
(** datatype mfix **) |
|
113 |
||
| 62753 | 114 |
(*Mfix (sy, ty, c, ps, p, pos): |
| 62752 | 115 |
sy: rhs of production as symbolic text |
| 240 | 116 |
ty: type description of production |
117 |
c: head of parse tree |
|
118 |
ps: priorities of arguments in sy |
|
| 62753 | 119 |
p: priority of production |
120 |
pos: source position*) |
|
| 240 | 121 |
|
| 62753 | 122 |
datatype mfix = Mfix of Symbol_Pos.T list * typ * string * int list * int * Position.T; |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
123 |
|
| 240 | 124 |
|
125 |
(* typ_to_nonterm *) |
|
126 |
||
| 865 | 127 |
fun typ_to_nt _ (Type (c, _)) = c |
128 |
| typ_to_nt default _ = default; |
|
129 |
||
|
764
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
clasohm
parents:
624
diff
changeset
|
130 |
(*get nonterminal for rhs*) |
|
42293
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents:
42288
diff
changeset
|
131 |
val typ_to_nonterm = typ_to_nt "any"; |
| 240 | 132 |
|
|
764
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
clasohm
parents:
624
diff
changeset
|
133 |
(*get nonterminal for lhs*) |
|
42293
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents:
42288
diff
changeset
|
134 |
val typ_to_nonterm1 = typ_to_nt "logic"; |
| 240 | 135 |
|
136 |
||
| 19004 | 137 |
(* read mixfix annotations *) |
| 4050 | 138 |
|
139 |
local |
|
| 19004 | 140 |
|
| 62752 | 141 |
open Basic_Symbol_Pos; |
142 |
||
143 |
fun scan_one pred = Scan.one (pred o Symbol_Pos.symbol); |
|
144 |
fun scan_many pred = Scan.many (pred o Symbol_Pos.symbol); |
|
145 |
fun scan_many1 pred = Scan.many1 (pred o Symbol_Pos.symbol); |
|
146 |
||
|
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59841
diff
changeset
|
147 |
val is_meta = member (op =) ["(", ")", "/", "_", "\<index>"];
|
| 19004 | 148 |
|
149 |
val scan_delim_char = |
|
| 62752 | 150 |
$$ "'" |-- scan_one ((not o Symbol.is_blank) andf Symbol.not_eof) || |
151 |
scan_one ((not o is_meta) andf (not o Symbol.is_blank) andf Symbol.not_eof); |
|
| 19004 | 152 |
|
153 |
fun read_int ["0", "0"] = ~1 |
|
154 |
| read_int cs = #1 (Library.read_int cs); |
|
| 4050 | 155 |
|
| 19004 | 156 |
val scan_sym = |
157 |
$$ "_" >> K (Argument ("", 0)) ||
|
|
|
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59841
diff
changeset
|
158 |
$$ "\<index>" >> K index || |
| 62752 | 159 |
$$ "(" |-- scan_many Symbol.is_digit >> (Bg o read_int o map Symbol_Pos.symbol) ||
|
| 19004 | 160 |
$$ ")" >> K En || |
161 |
$$ "/" -- $$ "/" >> K (Brk ~1) || |
|
| 62752 | 162 |
$$ "/" |-- scan_many Symbol.is_blank >> (Brk o length) || |
163 |
scan_many1 Symbol.is_blank >> (Space o Symbol_Pos.content) || |
|
164 |
Scan.repeat1 scan_delim_char >> (Delim o Symbol_Pos.content); |
|
| 14819 | 165 |
|
| 19004 | 166 |
val scan_symb = |
| 62764 | 167 |
Scan.trace scan_sym >> |
168 |
(fn (syms, trace) => SOME (syms, Position.set_range (Symbol_Pos.range trace))) || |
|
| 62752 | 169 |
$$ "'" -- scan_one Symbol.is_blank >> K NONE; |
| 19004 | 170 |
|
| 19305 | 171 |
val scan_symbs = Scan.repeat scan_symb --| Scan.ahead (~$$ "'"); |
| 4050 | 172 |
|
| 19004 | 173 |
in |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
174 |
|
| 62764 | 175 |
fun read_mfix ss = |
176 |
let |
|
177 |
val xsymbs = map_filter I (the (Scan.read Symbol_Pos.stopper scan_symbs ss)); |
|
178 |
val _ = Position.reports (maps reports_of xsymbs); |
|
179 |
in xsymbs end; |
|
| 19004 | 180 |
|
| 62764 | 181 |
val mfix_args = length o filter (is_argument o #1) o read_mfix o map (apsnd (K Position.none)); |
| 62752 | 182 |
val mixfix_args = mfix_args o Input.source_explode; |
| 19004 | 183 |
|
| 35390 | 184 |
val escape = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode; |
| 19004 | 185 |
|
| 4050 | 186 |
end; |
187 |
||
188 |
||
| 240 | 189 |
(* mfix_to_xprod *) |
190 |
||
| 62753 | 191 |
fun mfix_to_xprod logical_types (Mfix (sy, typ, const, pris, pri, pos)) = |
| 240 | 192 |
let |
| 62772 | 193 |
val _ = Position.report pos Markup.language_mixfix; |
| 62764 | 194 |
val symbs0 = read_mfix sy; |
195 |
||
| 62762 | 196 |
fun err_in_mixfix msg = error (msg ^ " in mixfix annotation" ^ Position.here pos); |
| 240 | 197 |
|
| 62762 | 198 |
fun check_blocks [] pending bad = pending @ bad |
199 |
| check_blocks ((Bg _, pos) :: rest) pending bad = check_blocks rest (pos :: pending) bad |
|
200 |
| check_blocks ((En, pos) :: rest) [] bad = check_blocks rest [] (pos :: bad) |
|
201 |
| check_blocks ((En, _) :: rest) (_ :: pending) bad = check_blocks rest pending bad |
|
202 |
| check_blocks (_ :: rest) pending bad = check_blocks rest pending bad; |
|
| 240 | 203 |
|
204 |
fun add_args [] ty [] = ([], typ_to_nonterm1 ty) |
|
| 62762 | 205 |
| add_args [] _ _ = err_in_mixfix "Too many precedences" |
206 |
| add_args ((sym as (Argument ("index", _), _)) :: syms) ty ps =
|
|
207 |
add_args syms ty ps |>> cons sym |
|
208 |
| add_args ((Argument _, pos) :: syms) (Type ("fun", [ty, tys])) [] =
|
|
209 |
add_args syms tys [] |>> cons (Argument (typ_to_nonterm ty, 0), pos) |
|
210 |
| add_args ((Argument _, pos) :: syms) (Type ("fun", [ty, tys])) (p :: ps) =
|
|
211 |
add_args syms tys ps |>> cons (Argument (typ_to_nonterm ty, p), pos) |
|
212 |
| add_args ((Argument _, _) :: _) _ _ = |
|
213 |
err_in_mixfix "More arguments than in corresponding type" |
|
214 |
| add_args (sym :: syms) ty ps = add_args syms ty ps |>> cons sym; |
|
215 |
||
216 |
fun logical_args (a as (Argument (s, p))) = |
|
217 |
if s <> "prop" andalso member (op =) logical_types s then Argument ("logic", p) else a
|
|
218 |
| logical_args a = a; |
|
| 240 | 219 |
|
220 |
fun rem_pri (Argument (s, _)) = Argument (s, chain_pri) |
|
221 |
| rem_pri sym = sym; |
|
222 |
||
| 62764 | 223 |
val indexes = filter (is_index o #1) symbs0; |
| 62762 | 224 |
val _ = |
225 |
if length indexes <= 1 then () |
|
226 |
else error ("More than one index argument" ^ Position.here_list (map #2 indexes));
|
|
| 2364 | 227 |
|
| 62764 | 228 |
val args = map_filter (fn (arg as Argument _, _) => SOME arg | _ => NONE) symbs0; |
|
44470
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm
parents:
43329
diff
changeset
|
229 |
val (const', typ', syntax_consts, parse_rules) = |
|
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm
parents:
43329
diff
changeset
|
230 |
if not (exists is_index args) then (const, typ, NONE, NONE) |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
231 |
else |
|
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
232 |
let |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35390
diff
changeset
|
233 |
val indexed_const = |
|
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35390
diff
changeset
|
234 |
if const <> "" then const ^ "_indexed" |
| 62762 | 235 |
else err_in_mixfix "Missing constant name for indexed syntax"; |
| 14697 | 236 |
val rangeT = Term.range_type typ handle Match => |
| 62762 | 237 |
err_in_mixfix "Missing structure argument for indexed syntax"; |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
238 |
|
|
43329
84472e198515
tuned signature: Name.invent and Name.invent_names;
wenzelm
parents:
43323
diff
changeset
|
239 |
val xs = map Ast.Variable (Name.invent Name.context "xa" (length args - 1)); |
| 19012 | 240 |
val (xs1, xs2) = chop (find_index is_index args) xs; |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
241 |
val i = Ast.Variable "i"; |
| 62762 | 242 |
val lhs = |
243 |
Ast.mk_appl (Ast.Constant indexed_const) |
|
244 |
(xs1 @ [Ast.mk_appl (Ast.Constant "_index") [i]] @ xs2); |
|
| 14697 | 245 |
val rhs = Ast.mk_appl (Ast.Constant const) (i :: xs); |
|
44470
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm
parents:
43329
diff
changeset
|
246 |
in (indexed_const, rangeT, SOME (indexed_const, const), SOME (lhs, rhs)) end; |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
247 |
|
| 62764 | 248 |
val (symbs1, lhs) = add_args symbs0 typ' pris; |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
249 |
|
| 2364 | 250 |
val copy_prod = |
| 20675 | 251 |
(lhs = "prop" orelse lhs = "logic") |
| 2364 | 252 |
andalso const <> "" |
| 62764 | 253 |
andalso not (null symbs1) |
254 |
andalso not (exists (is_delim o #1) symbs1); |
|
| 2364 | 255 |
val lhs' = |
| 62764 | 256 |
if copy_prod orelse lhs = "prop" andalso map #1 symbs1 = [Argument ("prop'", 0)] then lhs
|
|
42293
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents:
42288
diff
changeset
|
257 |
else if lhs = "prop" then "prop'" |
| 62762 | 258 |
else if member (op =) logical_types lhs then "logic" |
| 2364 | 259 |
else lhs; |
| 62764 | 260 |
val symbs2 = map (apfst logical_args) symbs1; |
| 240 | 261 |
|
| 62762 | 262 |
val _ = |
263 |
(pri :: pris) |> List.app (fn p => |
|
264 |
if p >= 0 andalso p <= 1000 then () |
|
265 |
else err_in_mixfix ("Precedence " ^ string_of_int p ^ " out of range"));
|
|
266 |
val _ = |
|
| 62764 | 267 |
(case check_blocks symbs2 [] [] of |
| 62762 | 268 |
[] => () |
269 |
| bad => error ("Unbalanced block parentheses" ^ Position.here_list bad));
|
|
270 |
||
| 62764 | 271 |
val xprod = XProd (lhs', map #1 symbs2, const', pri); |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
272 |
val xprod' = |
| 62762 | 273 |
if Lexicon.is_terminal lhs' then |
274 |
err_in_mixfix ("Illegal use of terminal " ^ quote lhs' ^ " as nonterminal")
|
|
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
275 |
else if const <> "" then xprod |
| 62764 | 276 |
else if length (filter (is_argument o #1) symbs2) <> 1 then |
| 62762 | 277 |
err_in_mixfix "Copy production must have exactly one argument" |
| 62764 | 278 |
else if exists (is_terminal o #1) symbs2 then xprod |
279 |
else XProd (lhs', map (rem_pri o #1) symbs2, "", chain_pri); |
|
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
280 |
|
|
44470
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm
parents:
43329
diff
changeset
|
281 |
in (xprod', syntax_consts, parse_rules) end; |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
282 |
|
| 240 | 283 |
|
284 |
||
285 |
(** datatype syn_ext **) |
|
286 |
||
287 |
datatype syn_ext = |
|
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35429
diff
changeset
|
288 |
Syn_Ext of {
|
| 240 | 289 |
xprods: xprod list, |
|
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42297
diff
changeset
|
290 |
consts: (string * string) list, |
| 21772 | 291 |
parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list, |
| 1510 | 292 |
parse_rules: (Ast.ast * Ast.ast) list, |
| 21772 | 293 |
parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list, |
|
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
294 |
print_translation: (string * ((Proof.context -> typ -> term list -> term) * stamp)) list, |
| 1510 | 295 |
print_rules: (Ast.ast * Ast.ast) list, |
| 42268 | 296 |
print_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list}; |
| 240 | 297 |
|
298 |
||
299 |
(* syn_ext *) |
|
300 |
||
| 59841 | 301 |
fun syn_ext' logical_types mfixes consts trfuns (parse_rules, print_rules) = |
| 240 | 302 |
let |
303 |
val (parse_ast_translation, parse_translation, print_translation, |
|
304 |
print_ast_translation) = trfuns; |
|
305 |
||
| 59841 | 306 |
val xprod_results = map (mfix_to_xprod logical_types) mfixes; |
|
44470
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm
parents:
43329
diff
changeset
|
307 |
val xprods = map #1 xprod_results; |
|
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm
parents:
43329
diff
changeset
|
308 |
val consts' = map_filter #2 xprod_results; |
|
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm
parents:
43329
diff
changeset
|
309 |
val parse_rules' = rev (map_filter #3 xprod_results); |
|
42298
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
wenzelm
parents:
42297
diff
changeset
|
310 |
val mfix_consts = map (fn Mfix x => (#3 x, "")) mfixes @ map (fn XProd x => (#3 x, "")) xprods; |
| 240 | 311 |
in |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
35429
diff
changeset
|
312 |
Syn_Ext {
|
| 624 | 313 |
xprods = xprods, |
|
44470
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
wenzelm
parents:
43329
diff
changeset
|
314 |
consts = mfix_consts @ consts' @ consts, |
| 240 | 315 |
parse_ast_translation = parse_ast_translation, |
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
316 |
parse_rules = parse_rules' @ parse_rules, |
| 240 | 317 |
parse_translation = parse_translation, |
318 |
print_translation = print_translation, |
|
|
12513
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
wenzelm
parents:
11546
diff
changeset
|
319 |
print_rules = map swap parse_rules' @ print_rules, |
| 42268 | 320 |
print_ast_translation = print_ast_translation} |
| 240 | 321 |
end; |
322 |
||
| 2382 | 323 |
|
| 59841 | 324 |
val syn_ext = syn_ext' []; |
|
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
325 |
|
| 42268 | 326 |
fun syn_ext_rules rules = syn_ext [] [] ([], [], [], []) rules; |
| 52143 | 327 |
fun syn_ext_trfuns trfuns = syn_ext [] [] trfuns ([], []); |
| 16610 | 328 |
|
| 15754 | 329 |
fun stamp_trfun s (c, f) = (c, (f, s)); |
330 |
fun mk_trfun tr = stamp_trfun (stamp ()) tr; |
|
|
29318
6337d1cb2ba0
added numeral, which supercedes num, xnum, float;
wenzelm
parents:
28904
diff
changeset
|
331 |
fun eq_trfun ((_, s1: stamp), (_, s2)) = s1 = s2; |
| 15754 | 332 |
|
| 240 | 333 |
end; |