| author | sultana | 
| Wed, 07 Mar 2012 13:00:30 +0000 | |
| changeset 46824 | 1257c80988cd | 
| parent 44470 | 6c6c31ef6bb2 | 
| child 52143 | 36ffe23b25f8 | 
| permissions | -rw-r--r-- | 
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42268diff
changeset | 1 | (* Title: Pure/Syntax/syntax_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 | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42268diff
changeset | 4 | Syntax extension. | 
| 240 | 5 | *) | 
| 6 | ||
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42268diff
changeset | 7 | signature SYNTAX_EXT = | 
| 4050 | 8 | sig | 
| 6760 | 9 | val dddot_indexname: indexname | 
| 35351 
7425aece4ee3
allow general mixfix syntax for type constructors;
 wenzelm parents: 
35110diff
changeset | 10 | datatype mfix = Mfix of string * typ * string * int list * int | 
| 
7425aece4ee3
allow general mixfix syntax for type constructors;
 wenzelm parents: 
35110diff
changeset | 11 | val err_in_mfix: string -> mfix -> 'a | 
| 1510 | 12 | val typ_to_nonterm: typ -> string | 
| 13 | datatype xsymb = | |
| 14 | Delim of string | | |
| 15 | Argument of string * int | | |
| 16 | Space of string | | |
| 17 | Bg of int | Brk of int | En | |
| 18 | datatype xprod = XProd of string * xsymb list * string * int | |
| 19 | val chain_pri: int | |
| 4701 | 20 | val delims_of: xprod list -> string list list | 
| 1510 | 21 | datatype syn_ext = | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
35429diff
changeset | 22 |     Syn_Ext of {
 | 
| 1510 | 23 | xprods: xprod list, | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 24 | consts: (string * string) list, | 
| 21772 | 25 | parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list, | 
| 1510 | 26 | parse_rules: (Ast.ast * Ast.ast) list, | 
| 21772 | 27 | parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list, | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 28 | print_translation: (string * ((Proof.context -> typ -> term list -> term) * stamp)) list, | 
| 1510 | 29 | print_rules: (Ast.ast * Ast.ast) list, | 
| 42268 | 30 | print_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list} | 
| 19004 | 31 | val mfix_delims: string -> string list | 
| 4054 | 32 | val mfix_args: string -> int | 
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42268diff
changeset | 33 | val escape: string -> string | 
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 34 | val syn_ext': (string -> bool) -> mfix list -> | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 35 | (string * string) list -> (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * | 
| 21772 | 36 | (string * ((Proof.context -> term list -> term) * stamp)) list * | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 37 | (string * ((Proof.context -> typ -> term list -> term) * stamp)) list * | 
| 42268 | 38 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> | 
| 39 | (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: 
42297diff
changeset | 40 | val syn_ext: mfix list -> (string * string) list -> | 
| 21772 | 41 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * | 
| 42 | (string * ((Proof.context -> term list -> term) * stamp)) list * | |
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 43 | (string * ((Proof.context -> typ -> term list -> term) * stamp)) list * | 
| 42268 | 44 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> | 
| 45 | (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext | |
| 14903 | 46 | val syn_ext_rules: (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext | 
| 47 | val syn_ext_trfuns: | |
| 15754 | 48 | (string * ((Ast.ast list -> Ast.ast) * stamp)) list * | 
| 49 | (string * ((term list -> term) * stamp)) list * | |
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 50 | (string * ((typ -> term list -> term) * stamp)) list * | 
| 15754 | 51 | (string * ((Ast.ast list -> Ast.ast) * stamp)) list -> syn_ext | 
| 16610 | 52 | val syn_ext_advanced_trfuns: | 
| 21772 | 53 | (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list * | 
| 54 | (string * ((Proof.context -> term list -> term) * stamp)) list * | |
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 55 | (string * ((Proof.context -> typ -> term list -> term) * stamp)) list * | 
| 21772 | 56 | (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: 
42268diff
changeset | 57 |   val stamp_trfun: stamp -> string * 'a -> string * ('a * stamp)
 | 
| 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42268diff
changeset | 58 |   val mk_trfun: string * 'a -> string * ('a * stamp)
 | 
| 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42268diff
changeset | 59 |   val eq_trfun: ('a * stamp) * ('a * stamp) -> bool
 | 
| 4050 | 60 | end; | 
| 240 | 61 | |
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42268diff
changeset | 62 | structure Syntax_Ext: SYNTAX_EXT = | 
| 240 | 63 | struct | 
| 64 | ||
| 2694 | 65 | |
| 240 | 66 | (** misc definitions **) | 
| 67 | ||
| 7472 | 68 | val dddot_indexname = ("dddot", 0);
 | 
| 6760 | 69 | |
| 70 | ||
| 240 | 71 | (** datatype xprod **) | 
| 72 | ||
| 73 | (*Delim s: delimiter s | |
| 74 | Argument (s, p): nonterminal s requiring priority >= p, or valued token | |
| 75 | Space s: some white space for printing | |
| 76 | Bg, Brk, En: blocks and breaks for pretty printing*) | |
| 77 | ||
| 78 | datatype xsymb = | |
| 79 | Delim of string | | |
| 80 | Argument of string * int | | |
| 81 | Space of string | | |
| 82 | Bg of int | Brk of int | En; | |
| 83 | ||
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 84 | fun is_delim (Delim _) = true | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 85 | | is_delim _ = false; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 86 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 87 | fun is_terminal (Delim _) = true | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 88 | | is_terminal (Argument (s, _)) = Lexicon.is_terminal s | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 89 | | is_terminal _ = false; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 90 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 91 | fun is_argument (Argument _) = true | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 92 | | is_argument _ = false; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 93 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 94 | fun is_index (Argument ("index", _)) = true
 | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 95 | | is_index _ = false; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 96 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 97 | val index = Argument ("index", 1000);
 | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 98 | |
| 240 | 99 | |
| 100 | (*XProd (lhs, syms, c, p): | |
| 101 | lhs: name of nonterminal on the lhs of the production | |
| 102 | syms: list of symbols on the rhs of the production | |
| 103 | c: head of parse tree | |
| 104 | p: priority of this production*) | |
| 105 | ||
| 106 | datatype xprod = XProd of string * xsymb list * string * int; | |
| 107 | ||
| 108 | val chain_pri = ~1; (*dummy for chain productions*) | |
| 109 | ||
| 110 | fun delims_of xprods = | |
| 19004 | 111 | fold (fn XProd (_, xsymbs, _, _) => | 
| 112 | fold (fn Delim s => insert (op =) s | _ => I) xsymbs) xprods [] | |
| 113 | |> map Symbol.explode; | |
| 240 | 114 | |
| 115 | ||
| 116 | ||
| 117 | (** datatype mfix **) | |
| 118 | ||
| 119 | (*Mfix (sy, ty, c, ps, p): | |
| 120 | sy: rhs of production as symbolic string | |
| 121 | ty: type description of production | |
| 122 | c: head of parse tree | |
| 123 | ps: priorities of arguments in sy | |
| 124 | p: priority of production*) | |
| 125 | ||
| 126 | datatype mfix = Mfix of string * typ * string * int list * int; | |
| 127 | ||
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 128 | fun err_in_mfix msg (Mfix (sy, _, const, _, _)) = | 
| 18678 | 129 |   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 | 130 | |
| 240 | 131 | |
| 132 | (* typ_to_nonterm *) | |
| 133 | ||
| 865 | 134 | fun typ_to_nt _ (Type (c, _)) = c | 
| 135 | | typ_to_nt default _ = default; | |
| 136 | ||
| 764 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 137 | (*get nonterminal for rhs*) | 
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 138 | val typ_to_nonterm = typ_to_nt "any"; | 
| 240 | 139 | |
| 764 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 clasohm parents: 
624diff
changeset | 140 | (*get nonterminal for lhs*) | 
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 141 | val typ_to_nonterm1 = typ_to_nt "logic"; | 
| 240 | 142 | |
| 143 | ||
| 19004 | 144 | (* read mixfix annotations *) | 
| 4050 | 145 | |
| 146 | local | |
| 19004 | 147 | |
| 20675 | 148 | val is_meta = member (op =) ["(", ")", "/", "_", "\\<index>"];
 | 
| 19004 | 149 | |
| 150 | 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 | 151 | $$ "'" |-- 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 | 152 | Scan.one ((not o is_meta) andf (not o Symbol.is_blank) andf Symbol.is_regular); | 
| 19004 | 153 | |
| 154 | fun read_int ["0", "0"] = ~1 | |
| 155 | | read_int cs = #1 (Library.read_int cs); | |
| 4050 | 156 | |
| 19004 | 157 | val scan_sym = | 
| 158 |   $$ "_" >> K (Argument ("", 0)) ||
 | |
| 159 | $$ "\\<index>" >> K index || | |
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21772diff
changeset | 160 |   $$ "(" |-- Scan.many Symbol.is_digit >> (Bg o read_int) ||
 | 
| 19004 | 161 | $$ ")" >> K En || | 
| 162 | $$ "/" -- $$ "/" >> K (Brk ~1) || | |
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21772diff
changeset | 163 | $$ "/" |-- 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 | 164 | Scan.many1 Symbol.is_blank >> (Space o implode) || | 
| 19004 | 165 | Scan.repeat1 scan_delim_char >> (Delim o implode); | 
| 14819 | 166 | |
| 19004 | 167 | val scan_symb = | 
| 168 | scan_sym >> SOME || | |
| 169 | $$ "'" -- Scan.one Symbol.is_blank >> K NONE; | |
| 170 | ||
| 19305 | 171 | val scan_symbs = Scan.repeat scan_symb --| Scan.ahead (~$$ "'"); | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19305diff
changeset | 172 | val read_symbs = map_filter I o the o Scan.read Symbol.stopper scan_symbs; | 
| 4050 | 173 | |
| 19004 | 174 | fun unique_index xsymbs = | 
| 33317 | 175 | if length (filter is_index xsymbs) <= 1 then xsymbs | 
| 19004 | 176 | else error "Duplicate index arguments (\\<index>)"; | 
| 4050 | 177 | |
| 19004 | 178 | in | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 179 | |
| 19004 | 180 | val read_mfix = unique_index o read_symbs o Symbol.explode; | 
| 181 | ||
| 182 | fun mfix_delims sy = fold_rev (fn Delim s => cons s | _ => I) (read_mfix sy) []; | |
| 33317 | 183 | val mfix_args = length o filter is_argument o read_mfix; | 
| 19004 | 184 | |
| 35390 | 185 | val escape = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode; | 
| 19004 | 186 | |
| 4050 | 187 | end; | 
| 188 | ||
| 189 | ||
| 240 | 190 | (* mfix_to_xprod *) | 
| 191 | ||
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 192 | fun mfix_to_xprod is_logtype (mfix as Mfix (sy, typ, const, pris, pri)) = | 
| 240 | 193 | let | 
| 194 | fun check_pri p = | |
| 42297 
140f283266b7
discontinued Syntax.max_pri, which is not really a symbolic parameter;
 wenzelm parents: 
42294diff
changeset | 195 | if p >= 0 andalso p <= 1000 then () | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 196 |       else err_in_mfix ("Precedence out of range: " ^ string_of_int p) mfix;
 | 
| 240 | 197 | |
| 198 | fun blocks_ok [] 0 = true | |
| 199 | | blocks_ok [] _ = false | |
| 200 | | blocks_ok (Bg _ :: syms) n = blocks_ok syms (n + 1) | |
| 201 | | blocks_ok (En :: _) 0 = false | |
| 202 | | blocks_ok (En :: syms) n = blocks_ok syms (n - 1) | |
| 203 | | blocks_ok (_ :: syms) n = blocks_ok syms n; | |
| 204 | ||
| 205 | fun check_blocks syms = | |
| 206 | if blocks_ok syms 0 then () | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 207 | else err_in_mfix "Unbalanced block parentheses" mfix; | 
| 240 | 208 | |
| 209 | ||
| 210 | val cons_fst = apfst o cons; | |
| 211 | ||
| 212 | fun add_args [] ty [] = ([], typ_to_nonterm1 ty) | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 213 | | add_args [] _ _ = err_in_mfix "Too many precedences" mfix | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 214 |       | add_args ((arg as Argument ("index", _)) :: syms) ty ps =
 | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 215 | cons_fst arg (add_args syms ty ps) | 
| 240 | 216 |       | add_args (Argument _ :: syms) (Type ("fun", [ty, tys])) [] =
 | 
| 217 | cons_fst (Argument (typ_to_nonterm ty, 0)) (add_args syms tys []) | |
| 218 |       | add_args (Argument _ :: syms) (Type ("fun", [ty, tys])) (p :: ps) =
 | |
| 219 | cons_fst (Argument (typ_to_nonterm ty, p)) (add_args syms tys ps) | |
| 220 | | add_args (Argument _ :: _) _ _ = | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 221 | err_in_mfix "More arguments than in corresponding type" mfix | 
| 240 | 222 | | add_args (sym :: syms) ty ps = cons_fst sym (add_args syms ty ps); | 
| 223 | ||
| 224 | fun rem_pri (Argument (s, _)) = Argument (s, chain_pri) | |
| 225 | | rem_pri sym = sym; | |
| 226 | ||
| 32785 | 227 | fun logify_types (a as (Argument (s, p))) = | 
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 228 |           if s <> "prop" andalso is_logtype s then Argument ("logic", p) else a
 | 
| 32785 | 229 | | logify_types a = a; | 
| 240 | 230 | |
| 2364 | 231 | |
| 19004 | 232 | val raw_symbs = read_mfix sy handle ERROR msg => err_in_mfix msg mfix; | 
| 33317 | 233 | val args = filter (fn Argument _ => true | _ => false) raw_symbs; | 
| 44470 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 wenzelm parents: 
43329diff
changeset | 234 | 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: 
43329diff
changeset | 235 | if not (exists is_index args) then (const, typ, NONE, NONE) | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 236 | else | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 237 | let | 
| 35429 
afa8cf9e63d8
authentic syntax for classes and type constructors;
 wenzelm parents: 
35390diff
changeset | 238 | val indexed_const = | 
| 
afa8cf9e63d8
authentic syntax for classes and type constructors;
 wenzelm parents: 
35390diff
changeset | 239 | if const <> "" then const ^ "_indexed" | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 240 | else err_in_mfix "Missing constant name for indexed syntax" mfix; | 
| 14697 | 241 | val rangeT = Term.range_type typ handle Match => | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 242 | err_in_mfix "Missing structure argument for indexed syntax" mfix; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 243 | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
43323diff
changeset | 244 | val xs = map Ast.Variable (Name.invent Name.context "xa" (length args - 1)); | 
| 19012 | 245 | 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 | 246 | val i = Ast.Variable "i"; | 
| 14697 | 247 | val lhs = Ast.mk_appl (Ast.Constant indexed_const) | 
| 248 | (xs1 @ [Ast.mk_appl (Ast.Constant "_index") [i]] @ xs2); | |
| 249 | 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: 
43329diff
changeset | 250 | 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: 
11546diff
changeset | 251 | |
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 252 | val (symbs, lhs) = add_args raw_symbs typ' pris; | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 253 | |
| 2364 | 254 | val copy_prod = | 
| 20675 | 255 | (lhs = "prop" orelse lhs = "logic") | 
| 2364 | 256 | andalso const <> "" | 
| 257 | andalso not (null symbs) | |
| 258 | andalso not (exists is_delim symbs); | |
| 259 | val lhs' = | |
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 260 |       if copy_prod orelse lhs = "prop" andalso symbs = [Argument ("prop'", 0)] then lhs
 | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 261 | else if lhs = "prop" then "prop'" | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 262 | else if is_logtype lhs then "logic" | 
| 2364 | 263 | else lhs; | 
| 32785 | 264 | val symbs' = map logify_types symbs; | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 265 | val xprod = XProd (lhs', symbs', const', pri); | 
| 240 | 266 | |
| 15570 | 267 | 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 | 268 | val xprod' = | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 269 |       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 | 270 | else if const <> "" then xprod | 
| 33317 | 271 | else if length (filter is_argument symbs') <> 1 then | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 272 | 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 | 273 | else if exists is_terminal symbs' then xprod | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 274 | else XProd (lhs', map rem_pri symbs', "", chain_pri); | 
| 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 275 | |
| 44470 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 wenzelm parents: 
43329diff
changeset | 276 | in (xprod', syntax_consts, parse_rules) end; | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 277 | |
| 240 | 278 | |
| 279 | ||
| 280 | (** datatype syn_ext **) | |
| 281 | ||
| 282 | datatype syn_ext = | |
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
35429diff
changeset | 283 |   Syn_Ext of {
 | 
| 240 | 284 | xprods: xprod list, | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 285 | consts: (string * string) list, | 
| 21772 | 286 | parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list, | 
| 1510 | 287 | parse_rules: (Ast.ast * Ast.ast) list, | 
| 21772 | 288 | parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list, | 
| 42247 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 wenzelm parents: 
42245diff
changeset | 289 | print_translation: (string * ((Proof.context -> typ -> term list -> term) * stamp)) list, | 
| 1510 | 290 | print_rules: (Ast.ast * Ast.ast) list, | 
| 42268 | 291 | print_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list}; | 
| 240 | 292 | |
| 293 | ||
| 294 | (* syn_ext *) | |
| 295 | ||
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 296 | fun syn_ext' is_logtype mfixes consts trfuns (parse_rules, print_rules) = | 
| 240 | 297 | let | 
| 298 | val (parse_ast_translation, parse_translation, print_translation, | |
| 299 | print_ast_translation) = trfuns; | |
| 300 | ||
| 44470 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 wenzelm parents: 
43329diff
changeset | 301 | val xprod_results = map (mfix_to_xprod is_logtype) mfixes; | 
| 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 wenzelm parents: 
43329diff
changeset | 302 | 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: 
43329diff
changeset | 303 | 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: 
43329diff
changeset | 304 | 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: 
42297diff
changeset | 305 | val mfix_consts = map (fn Mfix x => (#3 x, "")) mfixes @ map (fn XProd x => (#3 x, "")) xprods; | 
| 240 | 306 | in | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
35429diff
changeset | 307 |     Syn_Ext {
 | 
| 624 | 308 | 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: 
43329diff
changeset | 309 | consts = mfix_consts @ consts' @ consts, | 
| 240 | 310 | parse_ast_translation = parse_ast_translation, | 
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 311 | parse_rules = parse_rules' @ parse_rules, | 
| 240 | 312 | parse_translation = parse_translation, | 
| 313 | print_translation = print_translation, | |
| 12513 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 wenzelm parents: 
11546diff
changeset | 314 | print_rules = map swap parse_rules' @ print_rules, | 
| 42268 | 315 | print_ast_translation = print_ast_translation} | 
| 240 | 316 | end; | 
| 317 | ||
| 2382 | 318 | |
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 319 | val syn_ext = syn_ext' (K false); | 
| 555 
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
 wenzelm parents: 
441diff
changeset | 320 | |
| 42268 | 321 | fun syn_ext_rules rules = syn_ext [] [] ([], [], [], []) rules; | 
| 322 | fun syn_ext_advanced_trfuns trfuns = syn_ext [] [] trfuns ([], []); | |
| 2382 | 323 | |
| 16610 | 324 | fun syn_ext_trfuns (atrs, trs, tr's, atr's) = | 
| 325 | let fun simple (name, (f, s)) = (name, (K f, s)) in | |
| 326 | syn_ext_advanced_trfuns (map simple atrs, map simple trs, map simple tr's, map simple atr's) | |
| 327 | end; | |
| 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: 
28904diff
changeset | 331 | fun eq_trfun ((_, s1: stamp), (_, s2)) = s1 = s2; | 
| 15754 | 332 | |
| 240 | 333 | end; |