| author | haftmann | 
| Mon, 26 Oct 2009 15:15:59 +0100 | |
| changeset 33205 | 20587741a8d9 | 
| parent 33042 | ddf1f03a9ad9 | 
| child 33317 | b4534348b8fd | 
| permissions | -rw-r--r-- | 
| 240 | 1 | (* Title: Pure/Syntax/syn_ext.ML | 
| 911 
55754d6d399c
new in mixfix annotations: "' " (quote space) separates delimiters without
 wenzelm parents: 
865diff
changeset | 2 | Author: Markus Wenzel and Carsten Clasohm, TU Muenchen | 
| 240 | 3 | |
| 4 | Syntax extension (internal interface). | |
| 5 | *) | |
| 6 | ||
| 7 | signature SYN_EXT0 = | |
| 4050 | 8 | sig | 
| 6760 | 9 | val dddot_indexname: indexname | 
| 10 | val constrainC: string | |
| 240 | 11 | val typeT: typ | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
26704diff
changeset | 12 | val spropT: typ | 
| 5690 | 13 | val max_pri: int | 
| 15754 | 14 |   val stamp_trfun: stamp -> string * 'a -> string * ('a * stamp)
 | 
| 15 |   val mk_trfun: string * 'a -> string * ('a * stamp)
 | |
| 16 |   val eq_trfun: ('a * stamp) * ('a * stamp) -> bool
 | |
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
25075diff
changeset | 17 | val tokentrans_mode: | 
| 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
25075diff
changeset | 18 | string -> (string * (Proof.context -> string -> Pretty.T)) list -> | 
| 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
25075diff
changeset | 19 | (string * string * (Proof.context -> string -> Pretty.T)) list | 
| 15835 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 20 | val standard_token_classes: string list | 
| 4050 | 21 | end; | 
| 240 | 22 | |
| 23 | signature SYN_EXT = | |
| 4050 | 24 | sig | 
| 240 | 25 | include SYN_EXT0 | 
| 1510 | 26 | val logic: string | 
| 27 | val args: string | |
| 28 | val cargs: string | |
| 30189 
3633f560f4c3
discontinued experimental support for Alice -- too hard to maintain its many language incompatibilities, never really worked anyway;
 wenzelm parents: 
29565diff
changeset | 29 | val any: string | 
| 1510 | 30 | val sprop: string | 
| 31 | val typ_to_nonterm: typ -> string | |
| 32 | datatype xsymb = | |
| 33 | Delim of string | | |
| 34 | Argument of string * int | | |
| 35 | Space of string | | |
| 36 | Bg of int | Brk of int | En | |
| 37 | datatype xprod = XProd of string * xsymb list * string * int | |
| 38 | val chain_pri: int | |
| 4701 | 39 | val delims_of: xprod list -> string list list | 
| 1510 | 40 | datatype mfix = Mfix of string * typ * string * int list * int | 
| 41 | datatype syn_ext = | |
| 42 |     SynExt of {
 | |
| 43 | xprods: xprod list, | |
| 44 | consts: string list, | |
| 2913 | 45 | prmodes: string list, | 
| 21772 | 46 | parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list, | 
| 1510 | 47 | parse_rules: (Ast.ast * Ast.ast) list, | 
| 21772 | 48 | parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list, | 
| 18857 | 49 | print_translation: | 
| 21772 | 50 | (string * ((Proof.context -> bool -> typ -> term list -> term) * stamp)) list, | 
| 1510 | 51 | print_rules: (Ast.ast * Ast.ast) list, | 
| 21772 | 52 | print_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list, | 
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
25075diff
changeset | 53 | token_translation: (string * string * (Proof.context -> string -> Pretty.T)) list} | 
| 19004 | 54 | val mfix_delims: string -> string list | 
| 4054 | 55 | val mfix_args: string -> int | 
| 12865 | 56 | val escape_mfix: string -> string | 
| 14903 | 57 | val syn_ext': bool -> (string -> bool) -> mfix list -> | 
| 21772 | 58 | string list -> (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * | 
| 59 | (string * ((Proof.context -> term list -> term) * stamp)) list * | |
| 60 | (string * ((Proof.context -> bool -> typ -> term list -> term) * stamp)) list * | |
| 61 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list | |
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
25075diff
changeset | 62 | -> (string * string * (Proof.context -> string -> Pretty.T)) list | 
| 1510 | 63 | -> (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext | 
| 14903 | 64 | val syn_ext: mfix list -> string list -> | 
| 21772 | 65 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * | 
| 66 | (string * ((Proof.context -> term list -> term) * stamp)) list * | |
| 67 | (string * ((Proof.context -> bool -> typ -> term list -> term) * stamp)) list * | |
| 68 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list | |
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
25075diff
changeset | 69 | -> (string * string * (Proof.context -> string -> Pretty.T)) list | 
| 1510 | 70 | -> (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext | 
| 14903 | 71 | val syn_ext_const_names: string list -> syn_ext | 
| 72 | val syn_ext_rules: (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext | |
| 73 | val syn_ext_trfuns: | |
| 15754 | 74 | (string * ((Ast.ast list -> Ast.ast) * stamp)) list * | 
| 75 | (string * ((term list -> term) * stamp)) list * | |
| 76 | (string * ((bool -> typ -> term list -> term) * stamp)) list * | |
| 77 | (string * ((Ast.ast list -> Ast.ast) * stamp)) list -> syn_ext | |
| 16610 | 78 | val syn_ext_advanced_trfuns: | 
| 21772 | 79 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * | 
| 80 | (string * ((Proof.context -> term list -> term) * stamp)) list * | |
| 81 | (string * ((Proof.context -> bool -> typ -> term list -> term) * stamp)) list * | |
| 82 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> syn_ext | |
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
25075diff
changeset | 83 | val syn_ext_tokentrfuns: (string * string * (Proof.context -> string -> Pretty.T)) list -> syn_ext | 
| 15835 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 84 | val standard_token_markers: string list | 
| 1510 | 85 | val pure_ext: syn_ext | 
| 4050 | 86 | end; | 
| 240 | 87 | |
| 15754 | 88 | structure SynExt: SYN_EXT = | 
| 240 | 89 | struct | 
| 90 | ||
| 2694 | 91 | |
| 240 | 92 | (** misc definitions **) | 
| 93 | ||
| 7472 | 94 | val dddot_indexname = ("dddot", 0);
 | 
| 6760 | 95 | val constrainC = "_constrain"; | 
| 96 | ||
| 97 | ||
| 240 | 98 | (* syntactic categories *) | 
| 99 | ||
| 100 | val logic = "logic"; | |
| 101 | val logicT = Type (logic, []); | |
| 102 | ||
| 103 | val args = "args"; | |
| 1178 | 104 | val cargs = "cargs"; | 
| 240 | 105 | |
| 330 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 clasohm parents: 
240diff
changeset | 106 | val typeT = Type ("type", []);
 | 
| 240 | 107 | |
| 764 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 108 | val sprop = "#prop"; | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 109 | val spropT = Type (sprop, []); | 
| 240 | 110 | |
| 30189 
3633f560f4c3
discontinued experimental support for Alice -- too hard to maintain its many language incompatibilities, never really worked anyway;
 wenzelm parents: 
29565diff
changeset | 111 | val any = "any"; | 
| 
3633f560f4c3
discontinued experimental support for Alice -- too hard to maintain its many language incompatibilities, never really worked anyway;
 wenzelm parents: 
29565diff
changeset | 112 | val anyT = Type (any, []); | 
| 624 | 113 | |
| 780 | 114 | |
| 240 | 115 | |
| 116 | (** datatype xprod **) | |
| 117 | ||
| 118 | (*Delim s: delimiter s | |
| 119 | Argument (s, p): nonterminal s requiring priority >= p, or valued token | |
| 120 | Space s: some white space for printing | |
| 121 | Bg, Brk, En: blocks and breaks for pretty printing*) | |
| 122 | ||
| 123 | datatype xsymb = | |
| 124 | Delim of string | | |
| 125 | Argument of string * int | | |
| 126 | Space of string | | |
| 127 | Bg of int | Brk of int | En; | |
| 128 | ||
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 129 | fun is_delim (Delim _) = true | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 130 | | is_delim _ = false; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 131 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 132 | fun is_terminal (Delim _) = true | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 133 | | is_terminal (Argument (s, _)) = Lexicon.is_terminal s | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 134 | | is_terminal _ = false; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 135 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 136 | fun is_argument (Argument _) = true | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 137 | | is_argument _ = false; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 138 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 139 | fun is_index (Argument ("index", _)) = true
 | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 140 | | is_index _ = false; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 141 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 142 | val index = Argument ("index", 1000);
 | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 143 | |
| 240 | 144 | |
| 145 | (*XProd (lhs, syms, c, p): | |
| 146 | lhs: name of nonterminal on the lhs of the production | |
| 147 | syms: list of symbols on the rhs of the production | |
| 148 | c: head of parse tree | |
| 149 | p: priority of this production*) | |
| 150 | ||
| 151 | datatype xprod = XProd of string * xsymb list * string * int; | |
| 152 | ||
| 153 | val max_pri = 1000; (*maximum legal priority*) | |
| 154 | val chain_pri = ~1; (*dummy for chain productions*) | |
| 155 | ||
| 156 | fun delims_of xprods = | |
| 19004 | 157 | fold (fn XProd (_, xsymbs, _, _) => | 
| 158 | fold (fn Delim s => insert (op =) s | _ => I) xsymbs) xprods [] | |
| 159 | |> map Symbol.explode; | |
| 240 | 160 | |
| 161 | ||
| 162 | ||
| 163 | (** datatype mfix **) | |
| 164 | ||
| 165 | (*Mfix (sy, ty, c, ps, p): | |
| 166 | sy: rhs of production as symbolic string | |
| 167 | ty: type description of production | |
| 168 | c: head of parse tree | |
| 169 | ps: priorities of arguments in sy | |
| 170 | p: priority of production*) | |
| 171 | ||
| 172 | datatype mfix = Mfix of string * typ * string * int list * int; | |
| 173 | ||
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 174 | fun err_in_mfix msg (Mfix (sy, _, const, _, _)) = | 
| 18678 | 175 |   cat_error msg ("in mixfix annotation " ^ quote sy ^ " for " ^ quote const);
 | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 176 | |
| 240 | 177 | |
| 178 | (* typ_to_nonterm *) | |
| 179 | ||
| 865 | 180 | fun typ_to_nt _ (Type (c, _)) = c | 
| 181 | | typ_to_nt default _ = default; | |
| 182 | ||
| 764 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 183 | (*get nonterminal for rhs*) | 
| 30189 
3633f560f4c3
discontinued experimental support for Alice -- too hard to maintain its many language incompatibilities, never really worked anyway;
 wenzelm parents: 
29565diff
changeset | 184 | val typ_to_nonterm = typ_to_nt any; | 
| 240 | 185 | |
| 764 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 186 | (*get nonterminal for lhs*) | 
| 865 | 187 | val typ_to_nonterm1 = typ_to_nt logic; | 
| 240 | 188 | |
| 189 | ||
| 19004 | 190 | (* read mixfix annotations *) | 
| 4050 | 191 | |
| 192 | local | |
| 19004 | 193 | |
| 20675 | 194 | val is_meta = member (op =) ["(", ")", "/", "_", "\\<index>"];
 | 
| 19004 | 195 | |
| 196 | val scan_delim_char = | |
| 23784 
75e6b9dd5336
Symbol.not_eof/sync is superceded by Symbol.is_regular (rules out further control symbols);
 wenzelm parents: 
23660diff
changeset | 197 | $$ "'" |-- Scan.one ((not o Symbol.is_blank) andf Symbol.is_regular) || | 
| 
75e6b9dd5336
Symbol.not_eof/sync is superceded by Symbol.is_regular (rules out further control symbols);
 wenzelm parents: 
23660diff
changeset | 198 | Scan.one ((not o is_meta) andf (not o Symbol.is_blank) andf Symbol.is_regular); | 
| 19004 | 199 | |
| 200 | fun read_int ["0", "0"] = ~1 | |
| 201 | | read_int cs = #1 (Library.read_int cs); | |
| 4050 | 202 | |
| 19004 | 203 | val scan_sym = | 
| 204 |   $$ "_" >> K (Argument ("", 0)) ||
 | |
| 205 | $$ "\\<index>" >> K index || | |
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21772diff
changeset | 206 |   $$ "(" |-- Scan.many Symbol.is_digit >> (Bg o read_int) ||
 | 
| 19004 | 207 | $$ ")" >> K En || | 
| 208 | $$ "/" -- $$ "/" >> K (Brk ~1) || | |
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21772diff
changeset | 209 | $$ "/" |-- Scan.many Symbol.is_blank >> (Brk o length) || | 
| 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21772diff
changeset | 210 | Scan.many1 Symbol.is_blank >> (Space o implode) || | 
| 19004 | 211 | Scan.repeat1 scan_delim_char >> (Delim o implode); | 
| 14819 | 212 | |
| 19004 | 213 | val scan_symb = | 
| 214 | scan_sym >> SOME || | |
| 215 | $$ "'" -- Scan.one Symbol.is_blank >> K NONE; | |
| 216 | ||
| 19305 | 217 | val scan_symbs = Scan.repeat scan_symb --| Scan.ahead (~$$ "'"); | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19305diff
changeset | 218 | val read_symbs = map_filter I o the o Scan.read Symbol.stopper scan_symbs; | 
| 4050 | 219 | |
| 19004 | 220 | fun unique_index xsymbs = | 
| 221 | if length (List.filter is_index xsymbs) <= 1 then xsymbs | |
| 222 | else error "Duplicate index arguments (\\<index>)"; | |
| 4050 | 223 | |
| 19004 | 224 | in | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 225 | |
| 19004 | 226 | val read_mfix = unique_index o read_symbs o Symbol.explode; | 
| 227 | ||
| 228 | fun mfix_delims sy = fold_rev (fn Delim s => cons s | _ => I) (read_mfix sy) []; | |
| 229 | val mfix_args = length o List.filter is_argument o read_mfix; | |
| 230 | ||
| 231 | val escape_mfix = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode; | |
| 232 | ||
| 4050 | 233 | end; | 
| 234 | ||
| 235 | ||
| 240 | 236 | (* mfix_to_xprod *) | 
| 237 | ||
| 14903 | 238 | fun mfix_to_xprod convert is_logtype (mfix as Mfix (sy, typ, const, pris, pri)) = | 
| 240 | 239 | let | 
| 240 | fun check_pri p = | |
| 241 | if p >= 0 andalso p <= max_pri then () | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 242 |       else err_in_mfix ("Precedence out of range: " ^ string_of_int p) mfix;
 | 
| 240 | 243 | |
| 244 | fun blocks_ok [] 0 = true | |
| 245 | | blocks_ok [] _ = false | |
| 246 | | blocks_ok (Bg _ :: syms) n = blocks_ok syms (n + 1) | |
| 247 | | blocks_ok (En :: _) 0 = false | |
| 248 | | blocks_ok (En :: syms) n = blocks_ok syms (n - 1) | |
| 249 | | blocks_ok (_ :: syms) n = blocks_ok syms n; | |
| 250 | ||
| 251 | fun check_blocks syms = | |
| 252 | if blocks_ok syms 0 then () | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 253 | else err_in_mfix "Unbalanced block parentheses" mfix; | 
| 240 | 254 | |
| 255 | ||
| 256 | val cons_fst = apfst o cons; | |
| 257 | ||
| 258 | fun add_args [] ty [] = ([], typ_to_nonterm1 ty) | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 259 | | add_args [] _ _ = err_in_mfix "Too many precedences" mfix | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 260 |       | add_args ((arg as Argument ("index", _)) :: syms) ty ps =
 | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 261 | cons_fst arg (add_args syms ty ps) | 
| 240 | 262 |       | add_args (Argument _ :: syms) (Type ("fun", [ty, tys])) [] =
 | 
| 263 | cons_fst (Argument (typ_to_nonterm ty, 0)) (add_args syms tys []) | |
| 264 |       | add_args (Argument _ :: syms) (Type ("fun", [ty, tys])) (p :: ps) =
 | |
| 265 | cons_fst (Argument (typ_to_nonterm ty, p)) (add_args syms tys ps) | |
| 266 | | add_args (Argument _ :: _) _ _ = | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 267 | err_in_mfix "More arguments than in corresponding type" mfix | 
| 240 | 268 | | add_args (sym :: syms) ty ps = cons_fst sym (add_args syms ty ps); | 
| 269 | ||
| 270 | fun rem_pri (Argument (s, _)) = Argument (s, chain_pri) | |
| 271 | | rem_pri sym = sym; | |
| 272 | ||
| 32785 | 273 | fun logify_types (a as (Argument (s, p))) = | 
| 14903 | 274 | if s <> "prop" andalso is_logtype s then Argument (logic, p) else a | 
| 32785 | 275 | | logify_types a = a; | 
| 240 | 276 | |
| 2364 | 277 | |
| 19004 | 278 | val raw_symbs = read_mfix sy handle ERROR msg => err_in_mfix msg mfix; | 
| 15570 | 279 | val args = List.filter (fn Argument _ => true | _ => false) raw_symbs; | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 280 | val (const', typ', parse_rules) = | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 281 | if not (exists is_index args) then (const, typ, []) | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 282 | else | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 283 | let | 
| 14697 | 284 | val indexed_const = if const <> "" then "_indexed_" ^ const | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 285 | else err_in_mfix "Missing constant name for indexed syntax" mfix; | 
| 14697 | 286 | val rangeT = Term.range_type typ handle Match => | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 287 | err_in_mfix "Missing structure argument for indexed syntax" mfix; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 288 | |
| 20076 | 289 | val xs = map Ast.Variable (Name.invent_list [] "xa" (length args - 1)); | 
| 19012 | 290 | val (xs1, xs2) = chop (find_index is_index args) xs; | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 291 | val i = Ast.Variable "i"; | 
| 14697 | 292 | val lhs = Ast.mk_appl (Ast.Constant indexed_const) | 
| 293 | (xs1 @ [Ast.mk_appl (Ast.Constant "_index") [i]] @ xs2); | |
| 294 | val rhs = Ast.mk_appl (Ast.Constant const) (i :: xs); | |
| 295 | in (indexed_const, rangeT, [(lhs, rhs)]) end; | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 296 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 297 | val (symbs, lhs) = add_args raw_symbs typ' pris; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 298 | |
| 2364 | 299 | val copy_prod = | 
| 20675 | 300 | (lhs = "prop" orelse lhs = "logic") | 
| 2364 | 301 | andalso const <> "" | 
| 302 | andalso not (null symbs) | |
| 303 | andalso not (exists is_delim symbs); | |
| 304 | val lhs' = | |
| 305 | if convert andalso not copy_prod then | |
| 14903 | 306 | (if lhs = "prop" then sprop else if is_logtype lhs then logic else lhs) | 
| 2364 | 307 | else lhs; | 
| 32785 | 308 | val symbs' = map logify_types symbs; | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 309 | val xprod = XProd (lhs', symbs', const', pri); | 
| 240 | 310 | |
| 15570 | 311 | val _ = (List.app check_pri pris; check_pri pri; check_blocks symbs'); | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 312 | val xprod' = | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 313 |       if Lexicon.is_terminal lhs' then err_in_mfix ("Illegal lhs: " ^ lhs') mfix
 | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 314 | else if const <> "" then xprod | 
| 15570 | 315 | else if length (List.filter is_argument symbs') <> 1 then | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 316 | err_in_mfix "Copy production must have exactly one argument" mfix | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 317 | else if exists is_terminal symbs' then xprod | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 318 | else XProd (lhs', map rem_pri symbs', "", chain_pri); | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 319 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 320 | in (xprod', parse_rules) end; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 321 | |
| 240 | 322 | |
| 323 | ||
| 324 | (** datatype syn_ext **) | |
| 325 | ||
| 326 | datatype syn_ext = | |
| 327 |   SynExt of {
 | |
| 328 | xprods: xprod list, | |
| 329 | consts: string list, | |
| 2913 | 330 | prmodes: string list, | 
| 21772 | 331 | parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list, | 
| 1510 | 332 | parse_rules: (Ast.ast * Ast.ast) list, | 
| 21772 | 333 | parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list, | 
| 18857 | 334 | print_translation: | 
| 21772 | 335 | (string * ((Proof.context -> bool -> typ -> term list -> term) * stamp)) list, | 
| 1510 | 336 | print_rules: (Ast.ast * Ast.ast) list, | 
| 21772 | 337 | print_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list, | 
| 26704 
51ee753cc2e3
token translations: context dependent, result Pretty.T;
 wenzelm parents: 
25075diff
changeset | 338 | token_translation: (string * string * (Proof.context -> string -> Pretty.T)) list}; | 
| 240 | 339 | |
| 340 | ||
| 341 | (* syn_ext *) | |
| 342 | ||
| 14903 | 343 | fun syn_ext' convert is_logtype mfixes consts trfuns tokentrfuns (parse_rules, print_rules) = | 
| 240 | 344 | let | 
| 345 | val (parse_ast_translation, parse_translation, print_translation, | |
| 346 | print_ast_translation) = trfuns; | |
| 347 | ||
| 14903 | 348 | val (xprods, parse_rules') = map (mfix_to_xprod convert is_logtype) mfixes | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19305diff
changeset | 349 | |> split_list |> apsnd (rev o flat); | 
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
19012diff
changeset | 350 | val mfix_consts = | 
| 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
19012diff
changeset | 351 | distinct (op =) (map (fn Mfix x => #3 x) mfixes @ map (fn XProd x => #3 x) xprods); | 
| 240 | 352 | in | 
| 353 |     SynExt {
 | |
| 624 | 354 | xprods = xprods, | 
| 33042 | 355 | consts = union (op =) consts mfix_consts, | 
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
19012diff
changeset | 356 | prmodes = distinct (op =) (map (fn (m, _, _) => m) tokentrfuns), | 
| 240 | 357 | parse_ast_translation = parse_ast_translation, | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 358 | parse_rules = parse_rules' @ parse_rules, | 
| 240 | 359 | parse_translation = parse_translation, | 
| 360 | print_translation = print_translation, | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 361 | print_rules = map swap parse_rules' @ print_rules, | 
| 2694 | 362 | print_ast_translation = print_ast_translation, | 
| 363 | token_translation = tokentrfuns} | |
| 240 | 364 | end; | 
| 365 | ||
| 2382 | 366 | |
| 14903 | 367 | val syn_ext = syn_ext' true (K false); | 
| 555 
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
 wenzelm parents: 
441diff
changeset | 368 | |
| 14903 | 369 | fun syn_ext_const_names cs = syn_ext [] cs ([], [], [], []) [] ([], []); | 
| 370 | fun syn_ext_rules rules = syn_ext [] [] ([], [], [], []) [] rules; | |
| 16610 | 371 | fun syn_ext_advanced_trfuns trfuns = syn_ext [] [] trfuns [] ([], []); | 
| 14903 | 372 | fun syn_ext_tokentrfuns trfuns = syn_ext [] [] ([], [], [], []) trfuns ([], []); | 
| 2382 | 373 | |
| 16610 | 374 | fun syn_ext_trfuns (atrs, trs, tr's, atr's) = | 
| 375 | let fun simple (name, (f, s)) = (name, (K f, s)) in | |
| 376 | syn_ext_advanced_trfuns (map simple atrs, map simple trs, map simple tr's, map simple atr's) | |
| 377 | end; | |
| 378 | ||
| 15754 | 379 | fun stamp_trfun s (c, f) = (c, (f, s)); | 
| 380 | fun mk_trfun tr = stamp_trfun (stamp ()) tr; | |
| 29318 
6337d1cb2ba0
added numeral, which supercedes num, xnum, float;
 wenzelm parents: 
28904diff
changeset | 381 | fun eq_trfun ((_, s1: stamp), (_, s2)) = s1 = s2; | 
| 15754 | 382 | |
| 240 | 383 | |
| 15835 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 384 | (* token translations *) | 
| 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 385 | |
| 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 386 | fun tokentrans_mode m trs = map (fn (s, f) => (m, s, f)) trs; | 
| 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 387 | |
| 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 388 | val standard_token_classes = | 
| 29318 
6337d1cb2ba0
added numeral, which supercedes num, xnum, float;
 wenzelm parents: 
28904diff
changeset | 389 | ["class", "tfree", "tvar", "free", "bound", "var", "numeral", "inner_string"]; | 
| 15835 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 390 | |
| 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 391 | val standard_token_markers = map (fn s => "_" ^ s) standard_token_classes; | 
| 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 392 | |
| 
fdf678bec567
added tokentrans_mode, standard_token_classes, standard_token_markers (from token_trans.ML);
 wenzelm parents: 
15754diff
changeset | 393 | |
| 764 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 394 | (* pure_ext *) | 
| 240 | 395 | |
| 14903 | 396 | val pure_ext = syn_ext' false (K false) | 
| 764 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 397 |   [Mfix ("_", spropT --> propT, "", [0], 0),
 | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 398 |    Mfix ("_", logicT --> anyT, "", [0], 0),
 | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 399 |    Mfix ("_", spropT --> anyT, "", [0], 0),
 | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 400 |    Mfix ("'(_')", logicT --> logicT, "", [0], max_pri),
 | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 401 |    Mfix ("'(_')", spropT --> spropT, "", [0], max_pri),
 | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 402 |    Mfix ("_::_",  [logicT, typeT] ---> logicT, "_constrain", [4, 0], 3),
 | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 403 |    Mfix ("_::_",  [spropT, typeT] ---> spropT, "_constrain", [4, 0], 3)]
 | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 404 | [] | 
| 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 405 | ([], [], [], []) | 
| 2694 | 406 | [] | 
| 764 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 407 | ([], []); | 
| 240 | 408 | |
| 409 | end; |