|
1 theory Outer_Syntax |
|
2 imports Base Main |
|
3 begin |
|
4 |
|
5 chapter {* Outer syntax --- the theory language \label{ch:outer-syntax} *} |
|
6 |
|
7 text {* |
|
8 The rather generic framework of Isabelle/Isar syntax emerges from |
|
9 three main syntactic categories: \emph{commands} of the top-level |
|
10 Isar engine (covering theory and proof elements), \emph{methods} for |
|
11 general goal refinements (analogous to traditional ``tactics''), and |
|
12 \emph{attributes} for operations on facts (within a certain |
|
13 context). Subsequently we give a reference of basic syntactic |
|
14 entities underlying Isabelle/Isar syntax in a bottom-up manner. |
|
15 Concrete theory and proof language elements will be introduced later |
|
16 on. |
|
17 |
|
18 \medskip In order to get started with writing well-formed |
|
19 Isabelle/Isar documents, the most important aspect to be noted is |
|
20 the difference of \emph{inner} versus \emph{outer} syntax. Inner |
|
21 syntax is that of Isabelle types and terms of the logic, while outer |
|
22 syntax is that of Isabelle/Isar theory sources (specifications and |
|
23 proofs). As a general rule, inner syntax entities may occur only as |
|
24 \emph{atomic entities} within outer syntax. For example, the string |
|
25 @{verbatim "\"x + y\""} and identifier @{verbatim z} are legal term |
|
26 specifications within a theory, while @{verbatim "x + y"} without |
|
27 quotes is not. |
|
28 |
|
29 Printed theory documents usually omit quotes to gain readability |
|
30 (this is a matter of {\LaTeX} macro setup, say via @{verbatim |
|
31 "\\isabellestyle"}, see also \cite{isabelle-sys}). Experienced |
|
32 users of Isabelle/Isar may easily reconstruct the lost technical |
|
33 information, while mere readers need not care about quotes at all. |
|
34 |
|
35 \medskip Isabelle/Isar input may contain any number of input |
|
36 termination characters ``@{verbatim ";"}'' (semicolon) to separate |
|
37 commands explicitly. This is particularly useful in interactive |
|
38 shell sessions to make clear where the current command is intended |
|
39 to end. Otherwise, the interpreter loop will continue to issue a |
|
40 secondary prompt ``@{verbatim "#"}'' until an end-of-command is |
|
41 clearly recognized from the input syntax, e.g.\ encounter of the |
|
42 next command keyword. |
|
43 |
|
44 More advanced interfaces such as Proof~General \cite{proofgeneral} |
|
45 do not require explicit semicolons, the amount of input text is |
|
46 determined automatically by inspecting the present content of the |
|
47 Emacs text buffer. In the printed presentation of Isabelle/Isar |
|
48 documents semicolons are omitted altogether for readability. |
|
49 |
|
50 \begin{warn} |
|
51 Proof~General requires certain syntax classification tables in |
|
52 order to achieve properly synchronized interaction with the |
|
53 Isabelle/Isar process. These tables need to be consistent with |
|
54 the Isabelle version and particular logic image to be used in a |
|
55 running session (common object-logics may well change the outer |
|
56 syntax). The standard setup should work correctly with any of the |
|
57 ``official'' logic images derived from Isabelle/HOL (including |
|
58 HOLCF etc.). Users of alternative logics may need to tell |
|
59 Proof~General explicitly, e.g.\ by giving an option @{verbatim "-k ZF"} |
|
60 (in conjunction with @{verbatim "-l ZF"}, to specify the default |
|
61 logic image). Note that option @{verbatim "-L"} does both |
|
62 of this at the same time. |
|
63 \end{warn} |
|
64 *} |
|
65 |
|
66 |
|
67 section {* Lexical matters \label{sec:outer-lex} *} |
|
68 |
|
69 text {* The outer lexical syntax consists of three main categories of |
|
70 syntax tokens: |
|
71 |
|
72 \begin{enumerate} |
|
73 |
|
74 \item \emph{major keywords} --- the command names that are available |
|
75 in the present logic session; |
|
76 |
|
77 \item \emph{minor keywords} --- additional literal tokens required |
|
78 by the syntax of commands; |
|
79 |
|
80 \item \emph{named tokens} --- various categories of identifiers etc. |
|
81 |
|
82 \end{enumerate} |
|
83 |
|
84 Major keywords and minor keywords are guaranteed to be disjoint. |
|
85 This helps user-interfaces to determine the overall structure of a |
|
86 theory text, without knowing the full details of command syntax. |
|
87 Internally, there is some additional information about the kind of |
|
88 major keywords, which approximates the command type (theory command, |
|
89 proof command etc.). |
|
90 |
|
91 Keywords override named tokens. For example, the presence of a |
|
92 command called @{verbatim term} inhibits the identifier @{verbatim |
|
93 term}, but the string @{verbatim "\"term\""} can be used instead. |
|
94 By convention, the outer syntax always allows quoted strings in |
|
95 addition to identifiers, wherever a named entity is expected. |
|
96 |
|
97 When tokenizing a given input sequence, the lexer repeatedly takes |
|
98 the longest prefix of the input that forms a valid token. Spaces, |
|
99 tabs, newlines and formfeeds between tokens serve as explicit |
|
100 separators. |
|
101 |
|
102 \medskip The categories for named tokens are defined once and for |
|
103 all as follows. |
|
104 |
|
105 \begin{center} |
|
106 \begin{supertabular}{rcl} |
|
107 @{syntax_def ident} & = & @{text "letter quasiletter\<^sup>*"} \\ |
|
108 @{syntax_def longident} & = & @{text "ident("}@{verbatim "."}@{text "ident)\<^sup>+"} \\ |
|
109 @{syntax_def symident} & = & @{text "sym\<^sup>+ | "}@{verbatim "\\"}@{verbatim "<"}@{text ident}@{verbatim ">"} \\ |
|
110 @{syntax_def nat} & = & @{text "digit\<^sup>+"} \\ |
|
111 @{syntax_def float} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\ |
|
112 @{syntax_def var} & = & @{verbatim "?"}@{text "ident | "}@{verbatim "?"}@{text ident}@{verbatim "."}@{text nat} \\ |
|
113 @{syntax_def typefree} & = & @{verbatim "'"}@{text ident} \\ |
|
114 @{syntax_def typevar} & = & @{verbatim "?"}@{text "typefree | "}@{verbatim "?"}@{text typefree}@{verbatim "."}@{text nat} \\ |
|
115 @{syntax_def string} & = & @{verbatim "\""} @{text "\<dots>"} @{verbatim "\""} \\ |
|
116 @{syntax_def altstring} & = & @{verbatim "`"} @{text "\<dots>"} @{verbatim "`"} \\ |
|
117 @{syntax_def verbatim} & = & @{verbatim "{*"} @{text "\<dots>"} @{verbatim "*"}@{verbatim "}"} \\[1ex] |
|
118 |
|
119 @{text letter} & = & @{text "latin | "}@{verbatim "\\"}@{verbatim "<"}@{text latin}@{verbatim ">"}@{text " | "}@{verbatim "\\"}@{verbatim "<"}@{text "latin latin"}@{verbatim ">"}@{text " | greek |"} \\ |
|
120 & & @{verbatim "\<^isub>"}@{text " | "}@{verbatim "\<^isup>"} \\ |
|
121 @{text quasiletter} & = & @{text "letter | digit | "}@{verbatim "_"}@{text " | "}@{verbatim "'"} \\ |
|
122 @{text latin} & = & @{verbatim a}@{text " | \<dots> | "}@{verbatim z}@{text " | "}@{verbatim A}@{text " | \<dots> | "}@{verbatim Z} \\ |
|
123 @{text digit} & = & @{verbatim "0"}@{text " | \<dots> | "}@{verbatim "9"} \\ |
|
124 @{text sym} & = & @{verbatim "!"}@{text " | "}@{verbatim "#"}@{text " | "}@{verbatim "$"}@{text " | "}@{verbatim "%"}@{text " | "}@{verbatim "&"}@{text " | "}@{verbatim "*"}@{text " | "}@{verbatim "+"}@{text " | "}@{verbatim "-"}@{text " | "}@{verbatim "/"}@{text " |"} \\ |
|
125 & & @{verbatim "<"}@{text " | "}@{verbatim "="}@{text " | "}@{verbatim ">"}@{text " | "}@{verbatim "?"}@{text " | "}@{verbatim "@"}@{text " | "}@{verbatim "^"}@{text " | "}@{verbatim "_"}@{text " | "}@{verbatim "|"}@{text " | "}@{verbatim "~"} \\ |
|
126 @{text greek} & = & @{verbatim "\<alpha>"}@{text " | "}@{verbatim "\<beta>"}@{text " | "}@{verbatim "\<gamma>"}@{text " | "}@{verbatim "\<delta>"}@{text " |"} \\ |
|
127 & & @{verbatim "\<epsilon>"}@{text " | "}@{verbatim "\<zeta>"}@{text " | "}@{verbatim "\<eta>"}@{text " | "}@{verbatim "\<theta>"}@{text " |"} \\ |
|
128 & & @{verbatim "\<iota>"}@{text " | "}@{verbatim "\<kappa>"}@{text " | "}@{verbatim "\<mu>"}@{text " | "}@{verbatim "\<nu>"}@{text " |"} \\ |
|
129 & & @{verbatim "\<xi>"}@{text " | "}@{verbatim "\<pi>"}@{text " | "}@{verbatim "\<rho>"}@{text " | "}@{verbatim "\<sigma>"}@{text " | "}@{verbatim "\<tau>"}@{text " |"} \\ |
|
130 & & @{verbatim "\<upsilon>"}@{text " | "}@{verbatim "\<phi>"}@{text " | "}@{verbatim "\<chi>"}@{text " | "}@{verbatim "\<psi>"}@{text " |"} \\ |
|
131 & & @{verbatim "\<omega>"}@{text " | "}@{verbatim "\<Gamma>"}@{text " | "}@{verbatim "\<Delta>"}@{text " | "}@{verbatim "\<Theta>"}@{text " |"} \\ |
|
132 & & @{verbatim "\<Lambda>"}@{text " | "}@{verbatim "\<Xi>"}@{text " | "}@{verbatim "\<Pi>"}@{text " | "}@{verbatim "\<Sigma>"}@{text " |"} \\ |
|
133 & & @{verbatim "\<Upsilon>"}@{text " | "}@{verbatim "\<Phi>"}@{text " | "}@{verbatim "\<Psi>"}@{text " | "}@{verbatim "\<Omega>"} \\ |
|
134 \end{supertabular} |
|
135 \end{center} |
|
136 |
|
137 A @{syntax_ref var} or @{syntax_ref typevar} describes an unknown, |
|
138 which is internally a pair of base name and index (ML type @{ML_type |
|
139 indexname}). These components are either separated by a dot as in |
|
140 @{text "?x.1"} or @{text "?x7.3"} or run together as in @{text |
|
141 "?x1"}. The latter form is possible if the base name does not end |
|
142 with digits. If the index is 0, it may be dropped altogether: |
|
143 @{text "?x"} and @{text "?x0"} and @{text "?x.0"} all refer to the |
|
144 same unknown, with basename @{text "x"} and index 0. |
|
145 |
|
146 The syntax of @{syntax_ref string} admits any characters, including |
|
147 newlines; ``@{verbatim "\""}'' (double-quote) and ``@{verbatim |
|
148 "\\"}'' (backslash) need to be escaped by a backslash; arbitrary |
|
149 character codes may be specified as ``@{verbatim "\\"}@{text ddd}'', |
|
150 with three decimal digits. Alternative strings according to |
|
151 @{syntax_ref altstring} are analogous, using single back-quotes |
|
152 instead. |
|
153 |
|
154 The body of @{syntax_ref verbatim} may consist of any text not |
|
155 containing ``@{verbatim "*"}@{verbatim "}"}''; this allows |
|
156 convenient inclusion of quotes without further escapes. There is no |
|
157 way to escape ``@{verbatim "*"}@{verbatim "}"}''. If the quoted |
|
158 text is {\LaTeX} source, one may usually add some blank or comment |
|
159 to avoid the critical character sequence. |
|
160 |
|
161 Source comments take the form @{verbatim "(*"}~@{text |
|
162 "\<dots>"}~@{verbatim "*)"} and may be nested, although the user-interface |
|
163 might prevent this. Note that this form indicates source comments |
|
164 only, which are stripped after lexical analysis of the input. The |
|
165 Isar syntax also provides proper \emph{document comments} that are |
|
166 considered as part of the text (see \secref{sec:comments}). |
|
167 |
|
168 Common mathematical symbols such as @{text \<forall>} are represented in |
|
169 Isabelle as @{verbatim \<forall>}. There are infinitely many Isabelle |
|
170 symbols like this, although proper presentation is left to front-end |
|
171 tools such as {\LaTeX}, Proof~General, or Isabelle/jEdit. A list of |
|
172 predefined Isabelle symbols that work well with these tools is given |
|
173 in \appref{app:symbols}. Note that @{verbatim "\<lambda>"} does not belong |
|
174 to the @{text letter} category, since it is already used differently |
|
175 in the Pure term language. *} |
|
176 |
|
177 |
|
178 section {* Common syntax entities *} |
|
179 |
|
180 text {* |
|
181 We now introduce several basic syntactic entities, such as names, |
|
182 terms, and theorem specifications, which are factored out of the |
|
183 actual Isar language elements to be described later. |
|
184 *} |
|
185 |
|
186 |
|
187 subsection {* Names *} |
|
188 |
|
189 text {* Entity @{syntax name} usually refers to any name of types, |
|
190 constants, theorems etc.\ that are to be \emph{declared} or |
|
191 \emph{defined} (so qualified identifiers are excluded here). Quoted |
|
192 strings provide an escape for non-identifier names or those ruled |
|
193 out by outer syntax keywords (e.g.\ quoted @{verbatim "\"let\""}). |
|
194 Already existing objects are usually referenced by @{syntax |
|
195 nameref}. |
|
196 |
|
197 @{rail " |
|
198 @{syntax_def name}: @{syntax ident} | @{syntax symident} | |
|
199 @{syntax string} | @{syntax nat} |
|
200 ; |
|
201 @{syntax_def parname}: '(' @{syntax name} ')' |
|
202 ; |
|
203 @{syntax_def nameref}: @{syntax name} | @{syntax longident} |
|
204 "} |
|
205 *} |
|
206 |
|
207 |
|
208 subsection {* Numbers *} |
|
209 |
|
210 text {* The outer lexical syntax (\secref{sec:outer-lex}) admits |
|
211 natural numbers and floating point numbers. These are combined as |
|
212 @{syntax int} and @{syntax real} as follows. |
|
213 |
|
214 @{rail " |
|
215 @{syntax_def int}: @{syntax nat} | '-' @{syntax nat} |
|
216 ; |
|
217 @{syntax_def real}: @{syntax float} | @{syntax int} |
|
218 "} |
|
219 |
|
220 Note that there is an overlap with the category @{syntax name}, |
|
221 which also includes @{syntax nat}. |
|
222 *} |
|
223 |
|
224 |
|
225 subsection {* Comments \label{sec:comments} *} |
|
226 |
|
227 text {* Large chunks of plain @{syntax text} are usually given |
|
228 @{syntax verbatim}, i.e.\ enclosed in @{verbatim "{"}@{verbatim |
|
229 "*"}~@{text "\<dots>"}~@{verbatim "*"}@{verbatim "}"}. For convenience, |
|
230 any of the smaller text units conforming to @{syntax nameref} are |
|
231 admitted as well. A marginal @{syntax comment} is of the form |
|
232 @{verbatim "--"}~@{syntax text}. Any number of these may occur |
|
233 within Isabelle/Isar commands. |
|
234 |
|
235 @{rail " |
|
236 @{syntax_def text}: @{syntax verbatim} | @{syntax nameref} |
|
237 ; |
|
238 @{syntax_def comment}: '--' @{syntax text} |
|
239 "} |
|
240 *} |
|
241 |
|
242 |
|
243 subsection {* Type classes, sorts and arities *} |
|
244 |
|
245 text {* |
|
246 Classes are specified by plain names. Sorts have a very simple |
|
247 inner syntax, which is either a single class name @{text c} or a |
|
248 list @{text "{c\<^sub>1, \<dots>, c\<^sub>n}"} referring to the |
|
249 intersection of these classes. The syntax of type arities is given |
|
250 directly at the outer level. |
|
251 |
|
252 @{rail " |
|
253 @{syntax_def classdecl}: @{syntax name} (('<' | '\<subseteq>') (@{syntax nameref} + ','))? |
|
254 ; |
|
255 @{syntax_def sort}: @{syntax nameref} |
|
256 ; |
|
257 @{syntax_def arity}: ('(' (@{syntax sort} + ',') ')')? @{syntax sort} |
|
258 "} |
|
259 *} |
|
260 |
|
261 |
|
262 subsection {* Types and terms \label{sec:types-terms} *} |
|
263 |
|
264 text {* |
|
265 The actual inner Isabelle syntax, that of types and terms of the |
|
266 logic, is far too sophisticated in order to be modelled explicitly |
|
267 at the outer theory level. Basically, any such entity has to be |
|
268 quoted to turn it into a single token (the parsing and type-checking |
|
269 is performed internally later). For convenience, a slightly more |
|
270 liberal convention is adopted: quotes may be omitted for any type or |
|
271 term that is already atomic at the outer level. For example, one |
|
272 may just write @{verbatim x} instead of quoted @{verbatim "\"x\""}. |
|
273 Note that symbolic identifiers (e.g.\ @{verbatim "++"} or @{text |
|
274 "\<forall>"} are available as well, provided these have not been superseded |
|
275 by commands or other keywords already (such as @{verbatim "="} or |
|
276 @{verbatim "+"}). |
|
277 |
|
278 @{rail " |
|
279 @{syntax_def type}: @{syntax nameref} | @{syntax typefree} | |
|
280 @{syntax typevar} |
|
281 ; |
|
282 @{syntax_def term}: @{syntax nameref} | @{syntax var} |
|
283 ; |
|
284 @{syntax_def prop}: @{syntax term} |
|
285 "} |
|
286 |
|
287 Positional instantiations are indicated by giving a sequence of |
|
288 terms, or the placeholder ``@{text _}'' (underscore), which means to |
|
289 skip a position. |
|
290 |
|
291 @{rail " |
|
292 @{syntax_def inst}: '_' | @{syntax term} |
|
293 ; |
|
294 @{syntax_def insts}: (@{syntax inst} *) |
|
295 "} |
|
296 |
|
297 Type declarations and definitions usually refer to @{syntax |
|
298 typespec} on the left-hand side. This models basic type constructor |
|
299 application at the outer syntax level. Note that only plain postfix |
|
300 notation is available here, but no infixes. |
|
301 |
|
302 @{rail " |
|
303 @{syntax_def typespec}: |
|
304 (() | @{syntax typefree} | '(' ( @{syntax typefree} + ',' ) ')') @{syntax name} |
|
305 ; |
|
306 @{syntax_def typespec_sorts}: |
|
307 (() | (@{syntax typefree} ('::' @{syntax sort})?) | |
|
308 '(' ( (@{syntax typefree} ('::' @{syntax sort})?) + ',' ) ')') @{syntax name} |
|
309 "} |
|
310 *} |
|
311 |
|
312 |
|
313 subsection {* Term patterns and declarations \label{sec:term-decls} *} |
|
314 |
|
315 text {* Wherever explicit propositions (or term fragments) occur in a |
|
316 proof text, casual binding of schematic term variables may be given |
|
317 specified via patterns of the form ``@{text "(\<IS> p\<^sub>1 \<dots> p\<^sub>n)"}''. |
|
318 This works both for @{syntax term} and @{syntax prop}. |
|
319 |
|
320 @{rail " |
|
321 @{syntax_def term_pat}: '(' (@'is' @{syntax term} +) ')' |
|
322 ; |
|
323 @{syntax_def prop_pat}: '(' (@'is' @{syntax prop} +) ')' |
|
324 "} |
|
325 |
|
326 \medskip Declarations of local variables @{text "x :: \<tau>"} and |
|
327 logical propositions @{text "a : \<phi>"} represent different views on |
|
328 the same principle of introducing a local scope. In practice, one |
|
329 may usually omit the typing of @{syntax vars} (due to |
|
330 type-inference), and the naming of propositions (due to implicit |
|
331 references of current facts). In any case, Isar proof elements |
|
332 usually admit to introduce multiple such items simultaneously. |
|
333 |
|
334 @{rail " |
|
335 @{syntax_def vars}: (@{syntax name} +) ('::' @{syntax type})? |
|
336 ; |
|
337 @{syntax_def props}: @{syntax thmdecl}? (@{syntax prop} @{syntax prop_pat}? +) |
|
338 "} |
|
339 |
|
340 The treatment of multiple declarations corresponds to the |
|
341 complementary focus of @{syntax vars} versus @{syntax props}. In |
|
342 ``@{text "x\<^sub>1 \<dots> x\<^sub>n :: \<tau>"}'' the typing refers to all variables, while |
|
343 in @{text "a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"} the naming refers to all propositions |
|
344 collectively. Isar language elements that refer to @{syntax vars} |
|
345 or @{syntax props} typically admit separate typings or namings via |
|
346 another level of iteration, with explicit @{keyword_ref "and"} |
|
347 separators; e.g.\ see @{command "fix"} and @{command "assume"} in |
|
348 \secref{sec:proof-context}. |
|
349 *} |
|
350 |
|
351 |
|
352 subsection {* Attributes and theorems \label{sec:syn-att} *} |
|
353 |
|
354 text {* Attributes have their own ``semi-inner'' syntax, in the sense |
|
355 that input conforming to @{syntax args} below is parsed by the |
|
356 attribute a second time. The attribute argument specifications may |
|
357 be any sequence of atomic entities (identifiers, strings etc.), or |
|
358 properly bracketed argument lists. Below @{syntax atom} refers to |
|
359 any atomic entity, including any @{syntax keyword} conforming to |
|
360 @{syntax symident}. |
|
361 |
|
362 @{rail " |
|
363 @{syntax_def atom}: @{syntax nameref} | @{syntax typefree} | |
|
364 @{syntax typevar} | @{syntax var} | @{syntax nat} | @{syntax float} | |
|
365 @{syntax keyword} |
|
366 ; |
|
367 arg: @{syntax atom} | '(' @{syntax args} ')' | '[' @{syntax args} ']' |
|
368 ; |
|
369 @{syntax_def args}: arg * |
|
370 ; |
|
371 @{syntax_def attributes}: '[' (@{syntax nameref} @{syntax args} * ',') ']' |
|
372 "} |
|
373 |
|
374 Theorem specifications come in several flavors: @{syntax axmdecl} |
|
375 and @{syntax thmdecl} usually refer to axioms, assumptions or |
|
376 results of goal statements, while @{syntax thmdef} collects lists of |
|
377 existing theorems. Existing theorems are given by @{syntax thmref} |
|
378 and @{syntax thmrefs}, the former requires an actual singleton |
|
379 result. |
|
380 |
|
381 There are three forms of theorem references: |
|
382 \begin{enumerate} |
|
383 |
|
384 \item named facts @{text "a"}, |
|
385 |
|
386 \item selections from named facts @{text "a(i)"} or @{text "a(j - k)"}, |
|
387 |
|
388 \item literal fact propositions using @{syntax_ref altstring} syntax |
|
389 @{verbatim "`"}@{text "\<phi>"}@{verbatim "`"} (see also method |
|
390 @{method_ref fact}). |
|
391 |
|
392 \end{enumerate} |
|
393 |
|
394 Any kind of theorem specification may include lists of attributes |
|
395 both on the left and right hand sides; attributes are applied to any |
|
396 immediately preceding fact. If names are omitted, the theorems are |
|
397 not stored within the theorem database of the theory or proof |
|
398 context, but any given attributes are applied nonetheless. |
|
399 |
|
400 An extra pair of brackets around attributes (like ``@{text |
|
401 "[[simproc a]]"}'') abbreviates a theorem reference involving an |
|
402 internal dummy fact, which will be ignored later on. So only the |
|
403 effect of the attribute on the background context will persist. |
|
404 This form of in-place declarations is particularly useful with |
|
405 commands like @{command "declare"} and @{command "using"}. |
|
406 |
|
407 @{rail " |
|
408 @{syntax_def axmdecl}: @{syntax name} @{syntax attributes}? ':' |
|
409 ; |
|
410 @{syntax_def thmdecl}: thmbind ':' |
|
411 ; |
|
412 @{syntax_def thmdef}: thmbind '=' |
|
413 ; |
|
414 @{syntax_def thmref}: |
|
415 (@{syntax nameref} selection? | @{syntax altstring}) @{syntax attributes}? | |
|
416 '[' @{syntax attributes} ']' |
|
417 ; |
|
418 @{syntax_def thmrefs}: @{syntax thmref} + |
|
419 ; |
|
420 |
|
421 thmbind: @{syntax name} @{syntax attributes} | @{syntax name} | @{syntax attributes} |
|
422 ; |
|
423 selection: '(' ((@{syntax nat} | @{syntax nat} '-' @{syntax nat}?) + ',') ')' |
|
424 "} |
|
425 *} |
|
426 |
|
427 end |