author | wenzelm |
Sat, 28 Apr 2012 16:06:30 +0200 | |
changeset 47822 | 34b44d28fc4b |
parent 47114 | 7c9e31ffcd9e |
permissions | -rw-r--r-- |
27037 | 1 |
theory Outer_Syntax |
42651 | 2 |
imports Base Main |
27037 | 3 |
begin |
4 |
||
47114
7c9e31ffcd9e
updated theory header syntax and related details;
wenzelm
parents:
46282
diff
changeset
|
5 |
chapter {* Outer syntax --- the theory language \label{ch:outer-syntax} *} |
27037 | 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 |
||
28774 | 67 |
section {* Lexical matters \label{sec:outer-lex} *} |
27037 | 68 |
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
69 |
text {* The outer lexical syntax consists of three main categories of |
28776 | 70 |
syntax tokens: |
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
71 |
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
72 |
\begin{enumerate} |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
73 |
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
74 |
\item \emph{major keywords} --- the command names that are available |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
75 |
in the present logic session; |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
76 |
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
77 |
\item \emph{minor keywords} --- additional literal tokens required |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
78 |
by the syntax of commands; |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
79 |
|
28776 | 80 |
\item \emph{named tokens} --- various categories of identifiers etc. |
27037 | 81 |
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
82 |
\end{enumerate} |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
83 |
|
28776 | 84 |
Major keywords and minor keywords are guaranteed to be disjoint. |
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
85 |
This helps user-interfaces to determine the overall structure of a |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
86 |
theory text, without knowing the full details of command syntax. |
28776 | 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.). |
|
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
90 |
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
91 |
Keywords override named tokens. For example, the presence of a |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
92 |
command called @{verbatim term} inhibits the identifier @{verbatim |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
93 |
term}, but the string @{verbatim "\"term\""} can be used instead. |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
94 |
By convention, the outer syntax always allows quoted strings in |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
95 |
addition to identifiers, wherever a named entity is expected. |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
96 |
|
28776 | 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 |
||
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
102 |
\medskip The categories for named tokens are defined once and for |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
103 |
all as follows. |
27037 | 104 |
|
28776 | 105 |
\begin{center} |
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
106 |
\begin{supertabular}{rcl} |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
107 |
@{syntax_def ident} & = & @{text "letter quasiletter\<^sup>*"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
108 |
@{syntax_def longident} & = & @{text "ident("}@{verbatim "."}@{text "ident)\<^sup>+"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
109 |
@{syntax_def symident} & = & @{text "sym\<^sup>+ | "}@{verbatim "\\"}@{verbatim "<"}@{text ident}@{verbatim ">"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
110 |
@{syntax_def nat} & = & @{text "digit\<^sup>+"} \\ |
40290
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
wenzelm
parents:
35841
diff
changeset
|
111 |
@{syntax_def float} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\ |
28775
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
112 |
@{syntax_def var} & = & @{verbatim "?"}@{text "ident | "}@{verbatim "?"}@{text ident}@{verbatim "."}@{text nat} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
113 |
@{syntax_def typefree} & = & @{verbatim "'"}@{text ident} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
114 |
@{syntax_def typevar} & = & @{verbatim "?"}@{text "typefree | "}@{verbatim "?"}@{text typefree}@{verbatim "."}@{text nat} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
115 |
@{syntax_def string} & = & @{verbatim "\""} @{text "\<dots>"} @{verbatim "\""} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
116 |
@{syntax_def altstring} & = & @{verbatim "`"} @{text "\<dots>"} @{verbatim "`"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
117 |
@{syntax_def verbatim} & = & @{verbatim "{*"} @{text "\<dots>"} @{verbatim "*"}@{verbatim "}"} \\[1ex] |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
118 |
|
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
119 |
@{text letter} & = & @{text "latin | "}@{verbatim "\\"}@{verbatim "<"}@{text latin}@{verbatim ">"}@{text " | "}@{verbatim "\\"}@{verbatim "<"}@{text "latin latin"}@{verbatim ">"}@{text " | greek |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
120 |
& & @{verbatim "\<^isub>"}@{text " | "}@{verbatim "\<^isup>"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
121 |
@{text quasiletter} & = & @{text "letter | digit | "}@{verbatim "_"}@{text " | "}@{verbatim "'"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
122 |
@{text latin} & = & @{verbatim a}@{text " | \<dots> | "}@{verbatim z}@{text " | "}@{verbatim A}@{text " | \<dots> | "}@{verbatim Z} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
123 |
@{text digit} & = & @{verbatim "0"}@{text " | \<dots> | "}@{verbatim "9"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
124 |
@{text sym} & = & @{verbatim "!"}@{text " | "}@{verbatim "#"}@{text " | "}@{verbatim "$"}@{text " | "}@{verbatim "%"}@{text " | "}@{verbatim "&"}@{text " | "}@{verbatim "*"}@{text " | "}@{verbatim "+"}@{text " | "}@{verbatim "-"}@{text " | "}@{verbatim "/"}@{text " |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
125 |
& & @{verbatim "<"}@{text " | "}@{verbatim "="}@{text " | "}@{verbatim ">"}@{text " | "}@{verbatim "?"}@{text " | "}@{verbatim "@"}@{text " | "}@{verbatim "^"}@{text " | "}@{verbatim "_"}@{text " | "}@{verbatim "|"}@{text " | "}@{verbatim "~"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
126 |
@{text greek} & = & @{verbatim "\<alpha>"}@{text " | "}@{verbatim "\<beta>"}@{text " | "}@{verbatim "\<gamma>"}@{text " | "}@{verbatim "\<delta>"}@{text " |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
127 |
& & @{verbatim "\<epsilon>"}@{text " | "}@{verbatim "\<zeta>"}@{text " | "}@{verbatim "\<eta>"}@{text " | "}@{verbatim "\<theta>"}@{text " |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
128 |
& & @{verbatim "\<iota>"}@{text " | "}@{verbatim "\<kappa>"}@{text " | "}@{verbatim "\<mu>"}@{text " | "}@{verbatim "\<nu>"}@{text " |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
129 |
& & @{verbatim "\<xi>"}@{text " | "}@{verbatim "\<pi>"}@{text " | "}@{verbatim "\<rho>"}@{text " | "}@{verbatim "\<sigma>"}@{text " | "}@{verbatim "\<tau>"}@{text " |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
130 |
& & @{verbatim "\<upsilon>"}@{text " | "}@{verbatim "\<phi>"}@{text " | "}@{verbatim "\<chi>"}@{text " | "}@{verbatim "\<psi>"}@{text " |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
131 |
& & @{verbatim "\<omega>"}@{text " | "}@{verbatim "\<Gamma>"}@{text " | "}@{verbatim "\<Delta>"}@{text " | "}@{verbatim "\<Theta>"}@{text " |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
132 |
& & @{verbatim "\<Lambda>"}@{text " | "}@{verbatim "\<Xi>"}@{text " | "}@{verbatim "\<Pi>"}@{text " | "}@{verbatim "\<Sigma>"}@{text " |"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
133 |
& & @{verbatim "\<Upsilon>"}@{text " | "}@{verbatim "\<Phi>"}@{text " | "}@{verbatim "\<Psi>"}@{text " | "}@{verbatim "\<Omega>"} \\ |
d25fe9601dbd
tuned outer lexical syntax; fixed var/tvar: really need question marks here;
wenzelm
parents:
28774
diff
changeset
|
134 |
\end{supertabular} |
28776 | 135 |
\end{center} |
27037 | 136 |
|
28778 | 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 |
|
27037 | 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 |
|
28778 | 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 |
|
27037 | 155 |
containing ``@{verbatim "*"}@{verbatim "}"}''; this allows |
28778 | 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}). |
|
27037 | 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 |
|
47822
34b44d28fc4b
some updates concerning current Proof General 4.x, which lacks X-Symbol mode of 3.x;
wenzelm
parents:
47114
diff
changeset
|
171 |
tools such as {\LaTeX}, Proof~General, or Isabelle/jEdit. A list of |
34b44d28fc4b
some updates concerning current Proof General 4.x, which lacks X-Symbol mode of 3.x;
wenzelm
parents:
47114
diff
changeset
|
172 |
predefined Isabelle symbols that work well with these tools is given |
34b44d28fc4b
some updates concerning current Proof General 4.x, which lacks X-Symbol mode of 3.x;
wenzelm
parents:
47114
diff
changeset
|
173 |
in \appref{app:symbols}. Note that @{verbatim "\<lambda>"} does not belong |
34b44d28fc4b
some updates concerning current Proof General 4.x, which lacks X-Symbol mode of 3.x;
wenzelm
parents:
47114
diff
changeset
|
174 |
to the @{text letter} category, since it is already used differently |
34b44d28fc4b
some updates concerning current Proof General 4.x, which lacks X-Symbol mode of 3.x;
wenzelm
parents:
47114
diff
changeset
|
175 |
in the Pure term language. *} |
27037 | 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 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
189 |
text {* Entity @{syntax name} usually refers to any name of types, |
27037 | 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\""}). |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
194 |
Already existing objects are usually referenced by @{syntax |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
195 |
nameref}. |
27037 | 196 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
197 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
198 |
@{syntax_def name}: @{syntax ident} | @{syntax symident} | |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
199 |
@{syntax string} | @{syntax nat} |
27037 | 200 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
201 |
@{syntax_def parname}: '(' @{syntax name} ')' |
27037 | 202 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
203 |
@{syntax_def nameref}: @{syntax name} | @{syntax longident} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
204 |
"} |
40296 | 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 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
214 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
215 |
@{syntax_def int}: @{syntax nat} | '-' @{syntax nat} |
27037 | 216 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
217 |
@{syntax_def real}: @{syntax float} | @{syntax int} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
218 |
"} |
40296 | 219 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
220 |
Note that there is an overlap with the category @{syntax name}, |
40296 | 221 |
which also includes @{syntax nat}. |
27037 | 222 |
*} |
223 |
||
224 |
||
225 |
subsection {* Comments \label{sec:comments} *} |
|
226 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
227 |
text {* Large chunks of plain @{syntax text} are usually given |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
228 |
@{syntax verbatim}, i.e.\ enclosed in @{verbatim "{"}@{verbatim |
27037 | 229 |
"*"}~@{text "\<dots>"}~@{verbatim "*"}@{verbatim "}"}. For convenience, |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
230 |
any of the smaller text units conforming to @{syntax nameref} are |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
231 |
admitted as well. A marginal @{syntax comment} is of the form |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
232 |
@{verbatim "--"}~@{syntax text}. Any number of these may occur |
27037 | 233 |
within Isabelle/Isar commands. |
234 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
235 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
236 |
@{syntax_def text}: @{syntax verbatim} | @{syntax nameref} |
27037 | 237 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
238 |
@{syntax_def comment}: '--' @{syntax text} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
239 |
"} |
27037 | 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 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
252 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
253 |
@{syntax_def classdecl}: @{syntax name} (('<' | '\<subseteq>') (@{syntax nameref} + ','))? |
27037 | 254 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
255 |
@{syntax_def sort}: @{syntax nameref} |
27037 | 256 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
257 |
@{syntax_def arity}: ('(' (@{syntax sort} + ',') ')')? @{syntax sort} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
258 |
"} |
27037 | 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 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
278 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
279 |
@{syntax_def type}: @{syntax nameref} | @{syntax typefree} | |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
280 |
@{syntax typevar} |
27037 | 281 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
282 |
@{syntax_def term}: @{syntax nameref} | @{syntax var} |
27037 | 283 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
284 |
@{syntax_def prop}: @{syntax term} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
285 |
"} |
27037 | 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 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
291 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
292 |
@{syntax_def inst}: '_' | @{syntax term} |
27037 | 293 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
294 |
@{syntax_def insts}: (@{syntax inst} *) |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
295 |
"} |
27037 | 296 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
297 |
Type declarations and definitions usually refer to @{syntax |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
298 |
typespec} on the left-hand side. This models basic type constructor |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
299 |
application at the outer syntax level. Note that only plain postfix |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
300 |
notation is available here, but no infixes. |
27037 | 301 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
302 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
303 |
@{syntax_def typespec}: |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
304 |
(() | @{syntax typefree} | '(' ( @{syntax typefree} + ',' ) ')') @{syntax name} |
27037 | 305 |
; |
42705 | 306 |
@{syntax_def typespec_sorts}: |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
307 |
(() | (@{syntax typefree} ('::' @{syntax sort})?) | |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
308 |
'(' ( (@{syntax typefree} ('::' @{syntax sort})?) + ',' ) ')') @{syntax name} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
309 |
"} |
27037 | 310 |
*} |
311 |
||
312 |
||
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
313 |
subsection {* Term patterns and declarations \label{sec:term-decls} *} |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
314 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
315 |
text {* Wherever explicit propositions (or term fragments) occur in a |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
316 |
proof text, casual binding of schematic term variables may be given |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
317 |
specified via patterns of the form ``@{text "(\<IS> p\<^sub>1 \<dots> p\<^sub>n)"}''. |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
318 |
This works both for @{syntax term} and @{syntax prop}. |
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
319 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
320 |
@{rail " |
42705 | 321 |
@{syntax_def term_pat}: '(' (@'is' @{syntax term} +) ')' |
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
322 |
; |
42705 | 323 |
@{syntax_def prop_pat}: '(' (@'is' @{syntax prop} +) ')' |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
324 |
"} |
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
325 |
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
326 |
\medskip Declarations of local variables @{text "x :: \<tau>"} and |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
327 |
logical propositions @{text "a : \<phi>"} represent different views on |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
328 |
the same principle of introducing a local scope. In practice, one |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
329 |
may usually omit the typing of @{syntax vars} (due to |
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
330 |
type-inference), and the naming of propositions (due to implicit |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
331 |
references of current facts). In any case, Isar proof elements |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
332 |
usually admit to introduce multiple such items simultaneously. |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
333 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
334 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
335 |
@{syntax_def vars}: (@{syntax name} +) ('::' @{syntax type})? |
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
336 |
; |
42705 | 337 |
@{syntax_def props}: @{syntax thmdecl}? (@{syntax prop} @{syntax prop_pat}? +) |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
338 |
"} |
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
339 |
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
340 |
The treatment of multiple declarations corresponds to the |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
341 |
complementary focus of @{syntax vars} versus @{syntax props}. In |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
342 |
``@{text "x\<^sub>1 \<dots> x\<^sub>n :: \<tau>"}'' the typing refers to all variables, while |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
343 |
in @{text "a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"} the naming refers to all propositions |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
344 |
collectively. Isar language elements that refer to @{syntax vars} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
345 |
or @{syntax props} typically admit separate typings or namings via |
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
346 |
another level of iteration, with explicit @{keyword_ref "and"} |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
347 |
separators; e.g.\ see @{command "fix"} and @{command "assume"} in |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
348 |
\secref{sec:proof-context}. |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
349 |
*} |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
350 |
|
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
351 |
|
27037 | 352 |
subsection {* Attributes and theorems \label{sec:syn-att} *} |
353 |
||
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
354 |
text {* Attributes have their own ``semi-inner'' syntax, in the sense |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
355 |
that input conforming to @{syntax args} below is parsed by the |
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
356 |
attribute a second time. The attribute argument specifications may |
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
357 |
be any sequence of atomic entities (identifiers, strings etc.), or |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
358 |
properly bracketed argument lists. Below @{syntax atom} refers to |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
359 |
any atomic entity, including any @{syntax keyword} conforming to |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
360 |
@{syntax symident}. |
27037 | 361 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
362 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
363 |
@{syntax_def atom}: @{syntax nameref} | @{syntax typefree} | |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
364 |
@{syntax typevar} | @{syntax var} | @{syntax nat} | @{syntax float} | |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
365 |
@{syntax keyword} |
27037 | 366 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
367 |
arg: @{syntax atom} | '(' @{syntax args} ')' | '[' @{syntax args} ']' |
27037 | 368 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
369 |
@{syntax_def args}: arg * |
27037 | 370 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
371 |
@{syntax_def attributes}: '[' (@{syntax nameref} @{syntax args} * ',') ']' |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
372 |
"} |
27037 | 373 |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
374 |
Theorem specifications come in several flavors: @{syntax axmdecl} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
375 |
and @{syntax thmdecl} usually refer to axioms, assumptions or |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
376 |
results of goal statements, while @{syntax thmdef} collects lists of |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
377 |
existing theorems. Existing theorems are given by @{syntax thmref} |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
378 |
and @{syntax thmrefs}, the former requires an actual singleton |
27037 | 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 |
|
28754
6f2e67a3dfaa
moved section "Proof method expressions" to proof chapter;
wenzelm
parents:
28753
diff
changeset
|
390 |
@{method_ref fact}). |
27037 | 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 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
407 |
@{rail " |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
408 |
@{syntax_def axmdecl}: @{syntax name} @{syntax attributes}? ':' |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
409 |
; |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
410 |
@{syntax_def thmdecl}: thmbind ':' |
27037 | 411 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
412 |
@{syntax_def thmdef}: thmbind '=' |
27037 | 413 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
414 |
@{syntax_def thmref}: |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
415 |
(@{syntax nameref} selection? | @{syntax altstring}) @{syntax attributes}? | |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
416 |
'[' @{syntax attributes} ']' |
27037 | 417 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
418 |
@{syntax_def thmrefs}: @{syntax thmref} + |
27037 | 419 |
; |
420 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
421 |
thmbind: @{syntax name} @{syntax attributes} | @{syntax name} | @{syntax attributes} |
27037 | 422 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
423 |
selection: '(' ((@{syntax nat} | @{syntax nat} '-' @{syntax nat}?) + ',') ')' |
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40296
diff
changeset
|
424 |
"} |
27037 | 425 |
*} |
426 |
||
427 |
end |