| author | wenzelm | 
| Tue, 10 Dec 2024 21:06:04 +0100 | |
| changeset 81572 | 693a95492008 | 
| parent 81231 | 808d940fa838 | 
| permissions | -rw-r--r-- | 
| 384 | 1 | (* Title: Pure/Syntax/mixfix.ML | 
| 81229 | 2 | Author: Tobias Nipkow, TU Muenchen | 
| 3 | Author: Makarius | |
| 384 | 4 | |
| 18719 
dca3ae4f6dd6
moved pure syntax to Syntax/syntax.ML and pure_thy.ML;
 wenzelm parents: 
18673diff
changeset | 5 | Mixfix declarations, infixes, binders. | 
| 384 | 6 | *) | 
| 7 | ||
| 81230 | 8 | datatype mixfix = | 
| 9 | NoSyn | | |
| 10 | Mixfix of Input.source * int list * int * Position.range | | |
| 11 | Infix of Input.source * int * Position.range | | |
| 12 | Infixl of Input.source * int * Position.range | | |
| 13 | Infixr of Input.source * int * Position.range | | |
| 14 | Binder of Input.source * int * int * Position.range | | |
| 15 | Structure of Position.range; | |
| 384 | 16 | |
| 42287 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 wenzelm parents: 
42284diff
changeset | 17 | signature MIXFIX = | 
| 2199 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 wenzelm parents: 
1952diff
changeset | 18 | sig | 
| 81230 | 19 | datatype mixfix = datatype mixfix | 
| 62761 | 20 | val mixfix: string -> mixfix | 
| 62752 | 21 | val is_empty: mixfix -> bool | 
| 81228 | 22 | val is_infix: mixfix -> bool | 
| 62752 | 23 | val equal: mixfix * mixfix -> bool | 
| 24 | val range_of: mixfix -> Position.range | |
| 25 | val pos_of: mixfix -> Position.T | |
| 62759 | 26 | val reset_pos: mixfix -> mixfix | 
| 19271 | 27 | val pretty_mixfix: mixfix -> Pretty.T | 
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 28 | val mixfix_args: Proof.context -> mixfix -> int | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 29 | val default_constraint: Proof.context -> mixfix -> typ | 
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
35390diff
changeset | 30 | val make_type: int -> typ | 
| 42287 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 wenzelm parents: 
42284diff
changeset | 31 | val binder_name: string -> string | 
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 32 | val syn_ext_types: Proof.context -> (string * typ * mixfix) list -> Syntax_Ext.syn_ext | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 33 | val syn_ext_consts: Proof.context -> string list -> (string * typ * mixfix) list -> | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 34 | Syntax_Ext.syn_ext | 
| 2199 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 wenzelm parents: 
1952diff
changeset | 35 | end; | 
| 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 wenzelm parents: 
1952diff
changeset | 36 | |
| 15751 
65e4790c7914
identify binder translations only once (admits remove);
 wenzelm parents: 
15570diff
changeset | 37 | structure Mixfix: MIXFIX = | 
| 384 | 38 | struct | 
| 39 | ||
| 40 | (** mixfix declarations **) | |
| 41 | ||
| 81230 | 42 | datatype mixfix = datatype mixfix; | 
| 62752 | 43 | |
| 62761 | 44 | fun mixfix s = Mixfix (Input.string s, [], 1000, Position.no_range); | 
| 45 | ||
| 62752 | 46 | fun is_empty NoSyn = true | 
| 47 | | is_empty _ = false; | |
| 48 | ||
| 81228 | 49 | fun is_infix (Infix _) = true | 
| 50 | | is_infix (Infixl _) = true | |
| 51 | | is_infix (Infixr _) = true | |
| 52 | | is_infix _ = false; | |
| 53 | ||
| 62752 | 54 | fun equal (NoSyn, NoSyn) = true | 
| 55 | | equal (Mixfix (sy, ps, p, _), Mixfix (sy', ps', p', _)) = | |
| 56 | Input.equal_content (sy, sy') andalso ps = ps' andalso p = p' | |
| 57 | | equal (Infix (sy, p, _), Infix (sy', p', _)) = Input.equal_content (sy, sy') andalso p = p' | |
| 58 | | equal (Infixl (sy, p, _), Infixl (sy', p', _)) = Input.equal_content (sy, sy') andalso p = p' | |
| 59 | | equal (Infixr (sy, p, _), Infixr (sy', p', _)) = Input.equal_content (sy, sy') andalso p = p' | |
| 60 | | equal (Binder (sy, p, q, _), Binder (sy', p', q', _)) = | |
| 61 | Input.equal_content (sy, sy') andalso p = p' andalso q = q' | |
| 62 | | equal (Structure _, Structure _) = true | |
| 63 | | equal _ = false; | |
| 64 | ||
| 65 | fun range_of NoSyn = Position.no_range | |
| 66 | | range_of (Mixfix (_, _, _, range)) = range | |
| 67 | | range_of (Infix (_, _, range)) = range | |
| 68 | | range_of (Infixl (_, _, range)) = range | |
| 69 | | range_of (Infixr (_, _, range)) = range | |
| 70 | | range_of (Binder (_, _, _, range)) = range | |
| 71 | | range_of (Structure range) = range; | |
| 72 | ||
| 62797 | 73 | val pos_of = Position.range_position o range_of; | 
| 384 | 74 | |
| 62759 | 75 | fun reset_pos NoSyn = NoSyn | 
| 76 | | reset_pos (Mixfix (sy, ps, p, _)) = Mixfix (Input.reset_pos sy, ps, p, Position.no_range) | |
| 77 | | reset_pos (Infix (sy, p, _)) = Infix (Input.reset_pos sy, p, Position.no_range) | |
| 78 | | reset_pos (Infixl (sy, p, _)) = Infixl (Input.reset_pos sy, p, Position.no_range) | |
| 79 | | reset_pos (Infixr (sy, p, _)) = Infixr (Input.reset_pos sy, p, Position.no_range) | |
| 80 | | reset_pos (Binder (sy, p, q, _)) = Binder (Input.reset_pos sy, p, q, Position.no_range) | |
| 81 | | reset_pos (Structure _) = Structure Position.no_range; | |
| 82 | ||
| 384 | 83 | |
| 19271 | 84 | (* pretty_mixfix *) | 
| 85 | ||
| 86 | local | |
| 11920 | 87 | |
| 69583 | 88 | val template = Pretty.cartouche o Pretty.str o #1 o Input.source_content; | 
| 55763 | 89 | val keyword = Pretty.keyword2; | 
| 19271 | 90 | val parens = Pretty.enclose "(" ")";
 | 
| 91 | val brackets = Pretty.enclose "[" "]"; | |
| 92 | val int = Pretty.str o string_of_int; | |
| 93 | ||
| 94 | in | |
| 11920 | 95 | |
| 19271 | 96 | fun pretty_mixfix NoSyn = Pretty.str "" | 
| 62752 | 97 | | pretty_mixfix (Mixfix (s, ps, p, _)) = | 
| 68273 
53788963c4dc
pretty-print according to defaults of input syntax;
 wenzelm parents: 
68271diff
changeset | 98 | parens | 
| 
53788963c4dc
pretty-print according to defaults of input syntax;
 wenzelm parents: 
68271diff
changeset | 99 | (Pretty.breaks | 
| 69582 | 100 | (template s :: | 
| 68273 
53788963c4dc
pretty-print according to defaults of input syntax;
 wenzelm parents: 
68271diff
changeset | 101 | (if null ps then [] else [brackets (Pretty.commas (map int ps))]) @ | 
| 
53788963c4dc
pretty-print according to defaults of input syntax;
 wenzelm parents: 
68271diff
changeset | 102 | (if p = 1000 then [] else [int p]))) | 
| 69582 | 103 | | pretty_mixfix (Infix (s, p, _)) = parens (Pretty.breaks [keyword "infix", template s, int p]) | 
| 104 | | pretty_mixfix (Infixl (s, p, _)) = parens (Pretty.breaks [keyword "infixl", template s, int p]) | |
| 105 | | pretty_mixfix (Infixr (s, p, _)) = parens (Pretty.breaks [keyword "infixr", template s, int p]) | |
| 62752 | 106 | | pretty_mixfix (Binder (s, p1, p2, _)) = | 
| 68273 
53788963c4dc
pretty-print according to defaults of input syntax;
 wenzelm parents: 
68271diff
changeset | 107 | parens | 
| 
53788963c4dc
pretty-print according to defaults of input syntax;
 wenzelm parents: 
68271diff
changeset | 108 | (Pretty.breaks | 
| 69582 | 109 | ([keyword "binder", template s] @ | 
| 110 | (if p1 = p2 then [] else [brackets [int p1]]) @ [int p2])) | |
| 62752 | 111 | | pretty_mixfix (Structure _) = parens [keyword "structure"]; | 
| 19271 | 112 | |
| 113 | end; | |
| 11920 | 114 | |
| 115 | ||
| 12531 | 116 | (* syntax specifications *) | 
| 384 | 117 | |
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 118 | fun mixfix_args ctxt = | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 119 | fn NoSyn => 0 | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 120 | | Mixfix (sy, _, _, _) => Syntax_Ext.mixfix_args ctxt sy | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 121 | | Infix (sy, _, _) => 2 + Syntax_Ext.mixfix_args ctxt sy | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 122 | | Infixl (sy, _, _) => 2 + Syntax_Ext.mixfix_args ctxt sy | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 123 | | Infixr (sy, _, _) => 2 + Syntax_Ext.mixfix_args ctxt sy | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 124 | | Binder _ => 1 | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 125 | | Structure _ => 0; | 
| 4053 | 126 | |
| 127 | ||
| 62959 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 wenzelm parents: 
62797diff
changeset | 128 | (* default type constraint *) | 
| 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 wenzelm parents: 
62797diff
changeset | 129 | |
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 130 | fun default_constraint _ (Binder _) = (dummyT --> dummyT) --> dummyT | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 131 | | default_constraint ctxt mx = replicate (mixfix_args ctxt mx) dummyT ---> dummyT; | 
| 62959 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 wenzelm parents: 
62797diff
changeset | 132 | |
| 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 wenzelm parents: 
62797diff
changeset | 133 | |
| 69586 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 134 | (* mixfix template *) | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 135 | |
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 136 | fun mixfix_template (Mixfix (sy, _, _, _)) = SOME sy | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 137 | | mixfix_template (Infix (sy, _, _)) = SOME sy | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 138 | | mixfix_template (Infixl (sy, _, _)) = SOME sy | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 139 | | mixfix_template (Infixr (sy, _, _)) = SOME sy | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 140 | | mixfix_template (Binder (sy, _, _, _)) = SOME sy | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 141 | | mixfix_template _ = NONE; | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 142 | |
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 143 | fun mixfix_template_reports mx = | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 144 | if Options.default_bool "update_mixfix_cartouches" then | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 145 | (case mixfix_template mx of | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 146 | SOME sy => [((Input.pos_of sy, Markup.update), cartouche (#1 (Input.source_content sy)))] | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 147 | | NONE => []) | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 148 | else []; | 
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 149 | |
| 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 150 | |
| 80911 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 151 | (* infix notation *) | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 152 | |
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 153 | val prefix_bg = Symbol_Pos.explode0 "'(";
 | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 154 | val prefix_en = Symbol_Pos.explode0 "')"; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 155 | |
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 156 | val infix_bg = Symbol_Pos.explode0 "(";
 | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 157 | val infix_arg1 = Symbol_Pos.explode0 "_ "; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 158 | val infix_arg2 = Symbol_Pos.explode0 "/ _)"; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 159 | |
| 80920 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 160 | fun infix_block ctxt = | 
| 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 161 | Syntax_Ext.mfix_name ctxt #> | 
| 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 162 | Syntax_Ext.block_annotation 0 Markup_Kind.markup_infix #> | 
| 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 163 | Symbol_Pos.explode0 #> | 
| 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 164 | append infix_bg; | 
| 80911 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 165 | |
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 166 | |
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 167 | (* binder notation *) | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 168 | |
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 169 | val binder_stamp = stamp (); | 
| 81231 | 170 | |
| 171 | fun binder_name c = | |
| 172 | (if Lexicon.is_const c then Lexicon.unmark_const c | |
| 173 | else if Lexicon.is_fixed c then Lexicon.unmark_fixed c | |
| 174 | else c) |> Lexicon.mark_binder; | |
| 80911 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 175 | |
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 176 | val binder_bg = Symbol_Pos.explode0 "(";
 | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 177 | val binder_en = Symbol_Pos.explode0 "_./ _)"; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 178 | |
| 80920 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 179 | fun binder_block ctxt = | 
| 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 180 | Syntax_Ext.mfix_name ctxt #> | 
| 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 181 | Syntax_Ext.block_annotation 3 Markup_Kind.markup_binder #> | 
| 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 182 | Symbol_Pos.explode0 #> | 
| 
bbe2c56fe255
more uniform treatment of Markup.notation and Markup.expression: manage kinds via context;
 wenzelm parents: 
80911diff
changeset | 183 | append binder_bg; | 
| 80911 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 184 | |
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 185 | fun binder_typ _ (Type ("fun", [Type ("fun", [_, ty2]), ty3])) =
 | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 186 |       [Type ("idts", []), ty2] ---> ty3
 | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 187 |   | binder_typ c _ = error ("Bad type of binder: " ^ quote c);
 | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 188 | |
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 189 | |
| 384 | 190 | (* syn_ext_types *) | 
| 191 | ||
| 42293 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 192 | val typeT = Type ("type", []);
 | 
| 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 wenzelm parents: 
42288diff
changeset | 193 | fun make_type n = replicate n typeT ---> typeT; | 
| 35412 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 wenzelm parents: 
35390diff
changeset | 194 | |
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 195 | fun syn_ext_types ctxt type_decls = | 
| 384 | 196 | let | 
| 62796 | 197 | fun mk_infix sy ty t p1 p2 p3 pos = | 
| 80911 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 198 | let | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 199 | val ss = Input.source_explode sy; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 200 | val syms = infix_block ctxt ss @ infix_arg1 @ ss @ infix_arg2; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 201 | in Syntax_Ext.Mfix (syms, ty, t, [p1, p2], p3, pos) end; | 
| 384 | 202 | |
| 80906 | 203 | fun mfix_of (t, ty, mx) = | 
| 204 | (case mx of | |
| 205 | NoSyn => NONE | |
| 206 | | Mixfix (sy, ps, p, _) => | |
| 62796 | 207 | SOME (Syntax_Ext.Mfix (Input.source_explode sy, ty, t, ps, p, pos_of mx)) | 
| 80906 | 208 | | Infix (sy, p, _) => SOME (mk_infix sy ty t (p + 1) (p + 1) p (pos_of mx)) | 
| 209 | | Infixl (sy, p, _) => SOME (mk_infix sy ty t p (p + 1) p (pos_of mx)) | |
| 210 | | Infixr (sy, p, _) => SOME (mk_infix sy ty t (p + 1) p p (pos_of mx)) | |
| 211 |       | _ => error ("Bad mixfix declaration for " ^ quote t ^ Position.here (pos_of mx)));
 | |
| 384 | 212 | |
| 62753 | 213 | fun check_args (_, ty, mx) (SOME (Syntax_Ext.Mfix (sy, _, _, _, _, _))) = | 
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 214 | if length (Term.binder_types ty) = Syntax_Ext.mfix_args ctxt sy then () | 
| 62752 | 215 | else | 
| 216 |             error ("Bad number of type constructor arguments in mixfix annotation" ^
 | |
| 217 | Position.here (pos_of mx)) | |
| 35351 
7425aece4ee3
allow general mixfix syntax for type constructors;
 wenzelm parents: 
35130diff
changeset | 218 | | check_args _ NONE = (); | 
| 
7425aece4ee3
allow general mixfix syntax for type constructors;
 wenzelm parents: 
35130diff
changeset | 219 | |
| 80899 
51c338103975
proper Context_Position.report, following 5328d67ec647;
 wenzelm parents: 
80897diff
changeset | 220 | val _ = Context_Position.reports_text ctxt (maps (mixfix_template_reports o #3) type_decls); | 
| 69586 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 221 | |
| 35351 
7425aece4ee3
allow general mixfix syntax for type constructors;
 wenzelm parents: 
35130diff
changeset | 222 | val mfix = map mfix_of type_decls; | 
| 
7425aece4ee3
allow general mixfix syntax for type constructors;
 wenzelm parents: 
35130diff
changeset | 223 | val _ = map2 check_args type_decls mfix; | 
| 80748 
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
 wenzelm parents: 
69586diff
changeset | 224 | val consts = map (fn (t, _, _) => (t, [])) type_decls; | 
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 225 | in Syntax_Ext.syn_ext ctxt [] (map_filter I mfix) consts ([], [], [], []) ([], []) end; | 
| 384 | 226 | |
| 227 | ||
| 228 | (* syn_ext_consts *) | |
| 229 | ||
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 230 | fun syn_ext_consts ctxt logical_types const_decls = | 
| 384 | 231 | let | 
| 62796 | 232 | fun mk_infix sy ty c p1 p2 p3 pos = | 
| 80911 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 233 | let | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 234 | val ss0 = Input.source_explode (Input.reset_pos sy); | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 235 | val ss = Input.source_explode sy; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 236 | val syms = infix_block ctxt ss @ infix_arg1 @ ss @ infix_arg2; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 237 | in | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 238 | [Syntax_Ext.Mfix (prefix_bg @ ss0 @ prefix_en, ty, c, [], 1000, Position.none), | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 239 | Syntax_Ext.Mfix (syms, ty, c, [p1, p2], p3, pos)] | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 240 | end; | 
| 384 | 241 | |
| 80906 | 242 | fun mfix_of (c, ty, mx) = | 
| 243 | (case mx of | |
| 244 | NoSyn => [] | |
| 80911 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 245 | | Mixfix (sy, ps, p, _) => [Syntax_Ext.Mfix (Input.source_explode sy, ty, c, ps, p, pos_of mx)] | 
| 80906 | 246 | | Infix (sy, p, _) => mk_infix sy ty c (p + 1) (p + 1) p (pos_of mx) | 
| 247 | | Infixl (sy, p, _) => mk_infix sy ty c p (p + 1) p (pos_of mx) | |
| 248 | | Infixr (sy, p, _) => mk_infix sy ty c (p + 1) p p (pos_of mx) | |
| 249 | | Binder (sy, p, q, _) => | |
| 80911 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 250 | let | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 251 | val ss = Input.source_explode sy; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 252 | val syms = binder_block ctxt ss @ ss @ binder_en; | 
| 
8ad5e6df050b
block markup for specific notation, notably infix and binder;
 wenzelm parents: 
80906diff
changeset | 253 | in [Syntax_Ext.Mfix (syms, binder_typ c ty, binder_name c, [0, p], q, pos_of mx)] end | 
| 80906 | 254 |       | _ => error ("Bad mixfix declaration for " ^ quote c ^ Position.here (pos_of mx)));
 | 
| 384 | 255 | |
| 21534 
68f805e9db0b
Binder: syntax const is determined by binder_name, not its syntax;
 wenzelm parents: 
20892diff
changeset | 256 | fun binder (c, _, Binder _) = SOME (binder_name c, c) | 
| 15531 | 257 | | binder _ = NONE; | 
| 384 | 258 | |
| 80899 
51c338103975
proper Context_Position.report, following 5328d67ec647;
 wenzelm parents: 
80897diff
changeset | 259 | val _ = Context_Position.reports_text ctxt (maps (mixfix_template_reports o #3) const_decls); | 
| 69586 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 wenzelm parents: 
69584diff
changeset | 260 | |
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19467diff
changeset | 261 | val mfix = maps mfix_of const_decls; | 
| 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19467diff
changeset | 262 | val binders = map_filter binder const_decls; | 
| 35129 | 263 | val binder_trs = binders | 
| 52143 | 264 | |> map (Syntax_Ext.stamp_trfun binder_stamp o Syntax_Trans.mk_binder_tr); | 
| 35129 | 265 | val binder_trs' = binders | 
| 42288 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 wenzelm parents: 
42287diff
changeset | 266 | |> map (Syntax_Ext.stamp_trfun binder_stamp o | 
| 52143 | 267 | apsnd Syntax_Trans.non_typed_tr' o Syntax_Trans.mk_binder_tr' o swap); | 
| 42298 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 wenzelm parents: 
42297diff
changeset | 268 | |
| 80748 
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
 wenzelm parents: 
69586diff
changeset | 269 | val consts = | 
| 
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
 wenzelm parents: 
69586diff
changeset | 270 | map (fn (c, b) => (c, [b])) binders @ | 
| 
43c4817375bf
support for syntax const dependencies, with minimal integrity checks;
 wenzelm parents: 
69586diff
changeset | 271 | map (fn (c, _, _) => (c, [])) const_decls; | 
| 80897 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 272 | in | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 273 | Syntax_Ext.syn_ext ctxt logical_types mfix consts ([], binder_trs, binder_trs', []) ([], []) | 
| 
5328d67ec647
more explicit context for syn_ext/mixfix operations, but it often degenerates to background theory;
 wenzelm parents: 
80748diff
changeset | 274 | end; | 
| 384 | 275 | |
| 276 | end; |