| author | wenzelm | 
| Thu, 29 Apr 2010 17:47:53 +0200 | |
| changeset 36508 | 03d2a2d0ee4a | 
| parent 35414 | cc8e4276d093 | 
| child 36745 | 403585a89772 | 
| permissions | -rw-r--r-- | 
| 28762 | 1  | 
%  | 
2  | 
\begin{isabellebody}%
 | 
|
3  | 
\def\isabellecontext{Inner{\isacharunderscore}Syntax}%
 | 
|
4  | 
%  | 
|
5  | 
\isadelimtheory  | 
|
6  | 
%  | 
|
7  | 
\endisadelimtheory  | 
|
8  | 
%  | 
|
9  | 
\isatagtheory  | 
|
10  | 
\isacommand{theory}\isamarkupfalse%
 | 
|
11  | 
\ Inner{\isacharunderscore}Syntax\isanewline
 | 
|
12  | 
\isakeyword{imports}\ Main\isanewline
 | 
|
13  | 
\isakeyword{begin}%
 | 
|
14  | 
\endisatagtheory  | 
|
15  | 
{\isafoldtheory}%
 | 
|
16  | 
%  | 
|
17  | 
\isadelimtheory  | 
|
18  | 
%  | 
|
19  | 
\endisadelimtheory  | 
|
20  | 
%  | 
|
21  | 
\isamarkupchapter{Inner syntax --- the term language \label{ch:inner-syntax}%
 | 
|
22  | 
}  | 
|
23  | 
\isamarkuptrue%  | 
|
24  | 
%  | 
|
25  | 
\isamarkupsection{Printing logical entities%
 | 
|
26  | 
}  | 
|
27  | 
\isamarkuptrue%  | 
|
28  | 
%  | 
|
29  | 
\isamarkupsubsection{Diagnostic commands%
 | 
|
30  | 
}  | 
|
31  | 
\isamarkuptrue%  | 
|
32  | 
%  | 
|
33  | 
\begin{isamarkuptext}%
 | 
|
34  | 
\begin{matharray}{rcl}
 | 
|
35  | 
    \indexdef{}{command}{typ}\hypertarget{command.typ}{\hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
|
36  | 
    \indexdef{}{command}{term}\hypertarget{command.term}{\hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
|
37  | 
    \indexdef{}{command}{prop}\hypertarget{command.prop}{\hyperlink{command.prop}{\mbox{\isa{\isacommand{prop}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
|
38  | 
    \indexdef{}{command}{thm}\hypertarget{command.thm}{\hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
|
39  | 
    \indexdef{}{command}{prf}\hypertarget{command.prf}{\hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
|
40  | 
    \indexdef{}{command}{full\_prf}\hypertarget{command.full-prf}{\hyperlink{command.full-prf}{\mbox{\isa{\isacommand{full{\isacharunderscore}prf}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
|
41  | 
    \indexdef{}{command}{pr}\hypertarget{command.pr}{\hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}any\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
|
42  | 
  \end{matharray}
 | 
|
43  | 
||
44  | 
These diagnostic commands assist interactive development by printing  | 
|
45  | 
internal logical entities in a human-readable fashion.  | 
|
46  | 
||
47  | 
  \begin{rail}
 | 
|
48  | 
'typ' modes? type  | 
|
49  | 
;  | 
|
50  | 
'term' modes? term  | 
|
51  | 
;  | 
|
52  | 
'prop' modes? prop  | 
|
53  | 
;  | 
|
54  | 
'thm' modes? thmrefs  | 
|
55  | 
;  | 
|
56  | 
( 'prf' | 'full\_prf' ) modes? thmrefs?  | 
|
57  | 
;  | 
|
58  | 
    'pr' modes? nat? (',' nat)?
 | 
|
59  | 
;  | 
|
60  | 
||
61  | 
    modes: '(' (name + ) ')'
 | 
|
62  | 
;  | 
|
63  | 
  \end{rail}
 | 
|
64  | 
||
65  | 
  \begin{description}
 | 
|
66  | 
||
67  | 
  \item \hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}~\isa{{\isasymtau}} reads and prints types of the
 | 
|
68  | 
meta-logic according to the current theory or proof context.  | 
|
69  | 
||
70  | 
  \item \hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}~\isa{t} and \hyperlink{command.prop}{\mbox{\isa{\isacommand{prop}}}}~\isa{{\isasymphi}}
 | 
|
71  | 
read, type-check and print terms or propositions according to the  | 
|
72  | 
  current theory or proof context; the inferred type of \isa{t} is
 | 
|
73  | 
output as well. Note that these commands are also useful in  | 
|
74  | 
inspecting the current environment of term abbreviations.  | 
|
75  | 
||
76  | 
  \item \hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}}~\isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isachardoublequote}} retrieves
 | 
|
77  | 
theorems from the current theory or proof context. Note that any  | 
|
78  | 
attributes included in the theorem specifications are applied to a  | 
|
79  | 
temporary context derived from the current theory or proof; the  | 
|
80  | 
  result is discarded, i.e.\ attributes involved in \isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ a\isactrlsub n{\isachardoublequote}} do not have any permanent effect.
 | 
|
81  | 
||
82  | 
  \item \hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}} displays the (compact) proof term of the
 | 
|
83  | 
current proof state (if present), or of the given theorems. Note  | 
|
84  | 
that this requires proof terms to be switched on for the current  | 
|
85  | 
object logic (see the ``Proof terms'' section of the Isabelle  | 
|
86  | 
reference manual for information on how to do this).  | 
|
87  | 
||
88  | 
  \item \hyperlink{command.full-prf}{\mbox{\isa{\isacommand{full{\isacharunderscore}prf}}}} is like \hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}}, but displays
 | 
|
89  | 
the full proof term, i.e.\ also displays information omitted in the  | 
|
90  | 
  compact proof term, which is denoted by ``\isa{{\isacharunderscore}}'' placeholders
 | 
|
91  | 
there.  | 
|
92  | 
||
93  | 
  \item \hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}~\isa{{\isachardoublequote}goals{\isacharcomma}\ prems{\isachardoublequote}} prints the current
 | 
|
94  | 
proof state (if present), including the proof context, current facts  | 
|
95  | 
and goals. The optional limit arguments affect the number of goals  | 
|
96  | 
and premises to be displayed, which is initially 10 for both.  | 
|
97  | 
Omitting limit values leaves the current setting unchanged.  | 
|
98  | 
||
99  | 
  \end{description}
 | 
|
100  | 
||
101  | 
  All of the diagnostic commands above admit a list of \isa{modes}
 | 
|
102  | 
to be specified, which is appended to the current print mode (see  | 
|
103  | 
  also \cite{isabelle-ref}).  Thus the output behavior may be modified
 | 
|
104  | 
  according particular print mode features.  For example, \hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}~\isa{{\isachardoublequote}{\isacharparenleft}latex\ xsymbols{\isacharparenright}{\isachardoublequote}} would print the current proof state
 | 
|
105  | 
with mathematical symbols and special characters represented in  | 
|
106  | 
  {\LaTeX} source, according to the Isabelle style
 | 
|
107  | 
  \cite{isabelle-sys}.
 | 
|
108  | 
||
109  | 
  Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more
 | 
|
110  | 
systematic way to include formal items into the printed text  | 
|
111  | 
document.%  | 
|
112  | 
\end{isamarkuptext}%
 | 
|
113  | 
\isamarkuptrue%  | 
|
114  | 
%  | 
|
115  | 
\isamarkupsubsection{Details of printed content%
 | 
|
116  | 
}  | 
|
117  | 
\isamarkuptrue%  | 
|
118  | 
%  | 
|
119  | 
\begin{isamarkuptext}%
 | 
|
120  | 
\begin{mldecls} 
 | 
|
| 32836 | 121  | 
    \indexdef{}{ML}{show\_types}\verb|show_types: bool Unsynchronized.ref| & default \verb|false| \\
 | 
122  | 
    \indexdef{}{ML}{show\_sorts}\verb|show_sorts: bool Unsynchronized.ref| & default \verb|false| \\
 | 
|
123  | 
    \indexdef{}{ML}{show\_consts}\verb|show_consts: bool Unsynchronized.ref| & default \verb|false| \\
 | 
|
124  | 
    \indexdef{}{ML}{long\_names}\verb|long_names: bool Unsynchronized.ref| & default \verb|false| \\
 | 
|
125  | 
    \indexdef{}{ML}{short\_names}\verb|short_names: bool Unsynchronized.ref| & default \verb|false| \\
 | 
|
126  | 
    \indexdef{}{ML}{unique\_names}\verb|unique_names: bool Unsynchronized.ref| & default \verb|true| \\
 | 
|
127  | 
    \indexdef{}{ML}{show\_brackets}\verb|show_brackets: bool Unsynchronized.ref| & default \verb|false| \\
 | 
|
128  | 
    \indexdef{}{ML}{eta\_contract}\verb|eta_contract: bool Unsynchronized.ref| & default \verb|true| \\
 | 
|
129  | 
    \indexdef{}{ML}{goals\_limit}\verb|goals_limit: int Unsynchronized.ref| & default \verb|10| \\
 | 
|
130  | 
    \indexdef{}{ML}{Proof.show\_main\_goal}\verb|Proof.show_main_goal: bool Unsynchronized.ref| & default \verb|false| \\
 | 
|
131  | 
    \indexdef{}{ML}{show\_hyps}\verb|show_hyps: bool Unsynchronized.ref| & default \verb|false| \\
 | 
|
132  | 
    \indexdef{}{ML}{show\_tags}\verb|show_tags: bool Unsynchronized.ref| & default \verb|false| \\
 | 
|
133  | 
    \indexdef{}{ML}{show\_question\_marks}\verb|show_question_marks: bool Unsynchronized.ref| & default \verb|true| \\
 | 
|
| 28762 | 134  | 
  \end{mldecls}
 | 
135  | 
||
136  | 
These global ML variables control the detail of information that is  | 
|
137  | 
displayed for types, terms, theorems, goals etc.  | 
|
138  | 
||
139  | 
In interactive sessions, the user interface usually manages these  | 
|
140  | 
global parameters of the Isabelle process, even with some concept of  | 
|
141  | 
persistence. Nonetheless it is occasionally useful to manipulate ML  | 
|
142  | 
  variables directly, e.g.\ using \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} or \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}}.
 | 
|
143  | 
||
144  | 
Batch-mode logic sessions may be configured by putting appropriate  | 
|
145  | 
ML text directly into the \verb|ROOT.ML| file.  | 
|
146  | 
||
147  | 
  \begin{description}
 | 
|
148  | 
||
149  | 
\item \verb|show_types| and \verb|show_sorts| control printing of type  | 
|
150  | 
constraints for term variables, and sort constraints for type  | 
|
151  | 
variables. By default, neither of these are shown in output. If  | 
|
152  | 
\verb|show_sorts| is set to \verb|true|, types are always shown as  | 
|
153  | 
well.  | 
|
154  | 
||
155  | 
Note that displaying types and sorts may explain why a polymorphic  | 
|
156  | 
inference rule fails to resolve with some goal, or why a rewrite  | 
|
157  | 
rule does not apply as expected.  | 
|
158  | 
||
159  | 
\item \verb|show_consts| controls printing of types of constants when  | 
|
160  | 
displaying a goal state.  | 
|
161  | 
||
162  | 
Note that the output can be enormous, because polymorphic constants  | 
|
163  | 
often occur at several different type instances.  | 
|
164  | 
||
165  | 
\item \verb|long_names|, \verb|short_names|, and \verb|unique_names|  | 
|
166  | 
control the way of printing fully qualified internal names in  | 
|
167  | 
  external form.  See also \secref{sec:antiq} for the document
 | 
|
168  | 
antiquotation options of the same names.  | 
|
169  | 
||
170  | 
\item \verb|show_brackets| controls bracketing in pretty printed  | 
|
171  | 
output. If set to \verb|true|, all sub-expressions of the pretty  | 
|
172  | 
printing tree will be parenthesized, even if this produces malformed  | 
|
173  | 
term syntax! This crude way of showing the internal structure of  | 
|
174  | 
pretty printed entities may occasionally help to diagnose problems  | 
|
175  | 
with operator priorities, for example.  | 
|
176  | 
||
177  | 
  \item \verb|eta_contract| controls \isa{{\isachardoublequote}{\isasymeta}{\isachardoublequote}}-contracted printing of
 | 
|
178  | 
terms.  | 
|
179  | 
||
180  | 
  The \isa{{\isasymeta}}-contraction law asserts \isa{{\isachardoublequote}{\isacharparenleft}{\isasymlambda}x{\isachardot}\ f\ x{\isacharparenright}\ {\isasymequiv}\ f{\isachardoublequote}},
 | 
|
181  | 
  provided \isa{x} is not free in \isa{f}.  It asserts
 | 
|
182  | 
  \emph{extensionality} of functions: \isa{{\isachardoublequote}f\ {\isasymequiv}\ g{\isachardoublequote}} if \isa{{\isachardoublequote}f\ x\ {\isasymequiv}\ g\ x{\isachardoublequote}} for all \isa{x}.  Higher-order unification frequently puts
 | 
|
183  | 
  terms into a fully \isa{{\isasymeta}}-expanded form.  For example, if \isa{F} has type \isa{{\isachardoublequote}{\isacharparenleft}{\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}{\isacharparenright}\ {\isasymRightarrow}\ {\isasymtau}{\isachardoublequote}} then its expanded form is \isa{{\isachardoublequote}{\isasymlambda}h{\isachardot}\ F\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ h\ x{\isacharparenright}{\isachardoublequote}}.
 | 
|
184  | 
||
185  | 
  Setting \verb|eta_contract| makes Isabelle perform \isa{{\isasymeta}}-contractions before printing, so that \isa{{\isachardoublequote}{\isasymlambda}h{\isachardot}\ F\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ h\ x{\isacharparenright}{\isachardoublequote}}
 | 
|
186  | 
  appears simply as \isa{F}.
 | 
|
187  | 
||
188  | 
  Note that the distinction between a term and its \isa{{\isasymeta}}-expanded
 | 
|
189  | 
form occasionally matters. While higher-order resolution and  | 
|
190  | 
  rewriting operate modulo \isa{{\isachardoublequote}{\isasymalpha}{\isasymbeta}{\isasymeta}{\isachardoublequote}}-conversion, some other tools
 | 
|
191  | 
might look at terms more discretely.  | 
|
192  | 
||
193  | 
\item \verb|goals_limit| controls the maximum number of subgoals to  | 
|
194  | 
be shown in goal output.  | 
|
195  | 
||
196  | 
\item \verb|Proof.show_main_goal| controls whether the main result to  | 
|
197  | 
be proven should be displayed. This information might be relevant  | 
|
198  | 
for schematic goals, to inspect the current claim that has been  | 
|
199  | 
synthesized so far.  | 
|
200  | 
||
201  | 
\item \verb|show_hyps| controls printing of implicit hypotheses of  | 
|
202  | 
local facts. Normally, only those hypotheses are displayed that are  | 
|
203  | 
  \emph{not} covered by the assumptions of the current context: this
 | 
|
204  | 
situation indicates a fault in some tool being used.  | 
|
205  | 
||
206  | 
  By setting \verb|show_hyps| to \verb|true|, output of \emph{all}
 | 
|
207  | 
hypotheses can be enforced, which is occasionally useful for  | 
|
208  | 
diagnostic purposes.  | 
|
209  | 
||
210  | 
\item \verb|show_tags| controls printing of extra annotations within  | 
|
211  | 
theorems, such as internal position information, or the case names  | 
|
212  | 
  being attached by the attribute \hyperlink{attribute.case-names}{\mbox{\isa{case{\isacharunderscore}names}}}.
 | 
|
213  | 
||
214  | 
  Note that the \hyperlink{attribute.tagged}{\mbox{\isa{tagged}}} and \hyperlink{attribute.untagged}{\mbox{\isa{untagged}}}
 | 
|
215  | 
attributes provide low-level access to the collection of tags  | 
|
216  | 
associated with a theorem.  | 
|
217  | 
||
218  | 
\item \verb|show_question_marks| controls printing of question marks  | 
|
219  | 
  for schematic variables, such as \isa{{\isacharquery}x}.  Only the leading
 | 
|
220  | 
question mark is affected, the remaining text is unchanged  | 
|
221  | 
(including proper markup for schematic variables that might be  | 
|
222  | 
relevant for user interfaces).  | 
|
223  | 
||
224  | 
  \end{description}%
 | 
|
225  | 
\end{isamarkuptext}%
 | 
|
226  | 
\isamarkuptrue%  | 
|
227  | 
%  | 
|
228  | 
\isamarkupsubsection{Printing limits%
 | 
|
229  | 
}  | 
|
230  | 
\isamarkuptrue%  | 
|
231  | 
%  | 
|
232  | 
\begin{isamarkuptext}%
 | 
|
233  | 
\begin{mldecls}
 | 
|
| 30121 | 234  | 
    \indexdef{}{ML}{Pretty.setdepth}\verb|Pretty.setdepth: int -> unit| \\
 | 
235  | 
    \indexdef{}{ML}{Pretty.setmargin}\verb|Pretty.setmargin: int -> unit| \\
 | 
|
236  | 
    \indexdef{}{ML}{print\_depth}\verb|print_depth: int -> unit| \\
 | 
|
| 28762 | 237  | 
  \end{mldecls}
 | 
238  | 
||
239  | 
These ML functions set limits for pretty printed text.  | 
|
240  | 
||
241  | 
  \begin{description}
 | 
|
242  | 
||
243  | 
  \item \verb|Pretty.setdepth|~\isa{d} tells the pretty printer to
 | 
|
244  | 
  limit the printing depth to \isa{d}.  This affects the display of
 | 
|
245  | 
types, terms, theorems etc. The default value is 0, which permits  | 
|
246  | 
  printing to an arbitrary depth.  Other useful values for \isa{d}
 | 
|
247  | 
are 10 and 20.  | 
|
248  | 
||
249  | 
  \item \verb|Pretty.setmargin|~\isa{m} tells the pretty printer to
 | 
|
250  | 
  assume a right margin (page width) of \isa{m}.  The initial margin
 | 
|
251  | 
is 76, but user interfaces might adapt the margin automatically when  | 
|
252  | 
resizing windows.  | 
|
253  | 
||
254  | 
  \item \verb|print_depth|~\isa{n} limits the printing depth of the
 | 
|
255  | 
ML toplevel pretty printer; the precise effect depends on the ML  | 
|
256  | 
  compiler and run-time system.  Typically \isa{n} should be less
 | 
|
257  | 
than 10. Bigger values such as 100--1000 are useful for debugging.  | 
|
258  | 
||
259  | 
  \end{description}%
 | 
|
260  | 
\end{isamarkuptext}%
 | 
|
261  | 
\isamarkuptrue%  | 
|
262  | 
%  | 
|
263  | 
\isamarkupsection{Mixfix annotations%
 | 
|
264  | 
}  | 
|
265  | 
\isamarkuptrue%  | 
|
266  | 
%  | 
|
267  | 
\begin{isamarkuptext}%
 | 
|
268  | 
Mixfix annotations specify concrete \emph{inner syntax} of
 | 
|
| 
35351
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
32836 
diff
changeset
 | 
269  | 
Isabelle types and terms. Locally fixed parameters in toplevel  | 
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
32836 
diff
changeset
 | 
270  | 
theorem statements, locale specifications etc.\ also admit mixfix  | 
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
32836 
diff
changeset
 | 
271  | 
annotations.  | 
| 28762 | 272  | 
|
273  | 
  \indexouternonterm{infix}\indexouternonterm{mixfix}\indexouternonterm{structmixfix}
 | 
|
274  | 
  \begin{rail}
 | 
|
275  | 
    infix: '(' ('infix' | 'infixl' | 'infixr') string nat ')'
 | 
|
276  | 
;  | 
|
277  | 
    mixfix: infix | '(' string prios? nat? ')' | '(' 'binder' string prios? nat ')'
 | 
|
278  | 
;  | 
|
279  | 
    structmixfix: mixfix | '(' 'structure' ')'
 | 
|
280  | 
;  | 
|
281  | 
||
282  | 
prios: '[' (nat + ',') ']'  | 
|
283  | 
;  | 
|
284  | 
  \end{rail}
 | 
|
285  | 
||
286  | 
  Here the \railtok{string} specifications refer to the actual mixfix
 | 
|
287  | 
template, which may include literal text, spacing, blocks, and  | 
|
288  | 
  arguments (denoted by ``\isa{{\isacharunderscore}}''); the special symbol
 | 
|
289  | 
  ``\verb|\<index>|'' (printed as ``\isa{{\isachardoublequote}{\isasymindex}{\isachardoublequote}}'') represents an index
 | 
|
290  | 
argument that specifies an implicit structure reference (see also  | 
|
291  | 
  \secref{sec:locale}).  Infix and binder declarations provide common
 | 
|
292  | 
abbreviations for particular mixfix declarations. So in practice,  | 
|
293  | 
mixfix templates mostly degenerate to literal text for concrete  | 
|
294  | 
syntax, such as ``\verb|++|'' for an infix symbol.  | 
|
295  | 
||
296  | 
\medskip In full generality, mixfix declarations work as follows.  | 
|
297  | 
  Suppose a constant \isa{{\isachardoublequote}c\ {\isacharcolon}{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymRightarrow}\ {\isasymdots}\ {\isasymtau}\isactrlsub n\ {\isasymRightarrow}\ {\isasymtau}{\isachardoublequote}} is
 | 
|
298  | 
  annotated by \isa{{\isachardoublequote}{\isacharparenleft}mixfix\ {\isacharbrackleft}p\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ p\isactrlsub n{\isacharbrackright}\ p{\isacharparenright}{\isachardoublequote}}, where \isa{{\isachardoublequote}mixfix{\isachardoublequote}} is a string \isa{{\isachardoublequote}d\isactrlsub {\isadigit{0}}\ {\isacharunderscore}\ d\isactrlsub {\isadigit{1}}\ {\isacharunderscore}\ {\isasymdots}\ {\isacharunderscore}\ d\isactrlsub n{\isachardoublequote}} consisting of
 | 
|
299  | 
delimiters that surround argument positions as indicated by  | 
|
300  | 
underscores.  | 
|
301  | 
||
302  | 
Altogether this determines a production for a context-free priority  | 
|
303  | 
  grammar, where for each argument \isa{{\isachardoublequote}i{\isachardoublequote}} the syntactic category
 | 
|
304  | 
  is determined by \isa{{\isachardoublequote}{\isasymtau}\isactrlsub i{\isachardoublequote}} (with priority \isa{{\isachardoublequote}p\isactrlsub i{\isachardoublequote}}), and
 | 
|
305  | 
  the result category is determined from \isa{{\isachardoublequote}{\isasymtau}{\isachardoublequote}} (with
 | 
|
306  | 
  priority \isa{{\isachardoublequote}p{\isachardoublequote}}).  Priority specifications are optional, with
 | 
|
307  | 
default 0 for arguments and 1000 for the result.  | 
|
308  | 
||
309  | 
  Since \isa{{\isachardoublequote}{\isasymtau}{\isachardoublequote}} may be again a function type, the constant
 | 
|
310  | 
type scheme may have more argument positions than the mixfix  | 
|
311  | 
  pattern.  Printing a nested application \isa{{\isachardoublequote}c\ t\isactrlsub {\isadigit{1}}\ {\isasymdots}\ t\isactrlsub m{\isachardoublequote}} for
 | 
|
312  | 
  \isa{{\isachardoublequote}m\ {\isachargreater}\ n{\isachardoublequote}} works by attaching concrete notation only to the
 | 
|
313  | 
  innermost part, essentially by printing \isa{{\isachardoublequote}{\isacharparenleft}c\ t\isactrlsub {\isadigit{1}}\ {\isasymdots}\ t\isactrlsub n{\isacharparenright}\ {\isasymdots}\ t\isactrlsub m{\isachardoublequote}}
 | 
|
314  | 
instead. If a term has fewer arguments than specified in the mixfix  | 
|
315  | 
template, the concrete syntax is ignored.  | 
|
316  | 
||
317  | 
\medskip A mixfix template may also contain additional directives  | 
|
318  | 
for pretty printing, notably spaces, blocks, and breaks. The  | 
|
319  | 
general template format is a sequence over any of the following  | 
|
320  | 
entities.  | 
|
321  | 
||
322  | 
  \begin{description}
 | 
|
323  | 
||
324  | 
  \item \isa{{\isachardoublequote}d{\isachardoublequote}} is a delimiter, namely a non-empty sequence of
 | 
|
325  | 
characters other than the following special characters:  | 
|
326  | 
||
327  | 
\smallskip  | 
|
328  | 
  \begin{tabular}{ll}
 | 
|
329  | 
\verb|'| & single quote \\  | 
|
330  | 
\verb|_| & underscore \\  | 
|
331  | 
    \isa{{\isachardoublequote}{\isasymindex}{\isachardoublequote}} & index symbol \\
 | 
|
332  | 
\verb|(| & open parenthesis \\  | 
|
333  | 
\verb|)| & close parenthesis \\  | 
|
334  | 
\verb|/| & slash \\  | 
|
335  | 
  \end{tabular}
 | 
|
336  | 
\medskip  | 
|
337  | 
||
338  | 
\item \verb|'| escapes the special meaning of these  | 
|
339  | 
meta-characters, producing a literal version of the following  | 
|
340  | 
character, unless that is a blank.  | 
|
341  | 
||
342  | 
A single quote followed by a blank separates delimiters, without  | 
|
343  | 
affecting printing, but input tokens may have additional white space  | 
|
344  | 
here.  | 
|
345  | 
||
346  | 
\item \verb|_| is an argument position, which stands for a  | 
|
347  | 
certain syntactic category in the underlying grammar.  | 
|
348  | 
||
349  | 
  \item \isa{{\isachardoublequote}{\isasymindex}{\isachardoublequote}} is an indexed argument position; this is the place
 | 
|
350  | 
where implicit structure arguments can be attached.  | 
|
351  | 
||
352  | 
  \item \isa{{\isachardoublequote}s{\isachardoublequote}} is a non-empty sequence of spaces for printing.
 | 
|
353  | 
This and the following specifications do not affect parsing at all.  | 
|
354  | 
||
355  | 
  \item \verb|(|\isa{n} opens a pretty printing block.  The
 | 
|
356  | 
optional number specifies how much indentation to add when a line  | 
|
357  | 
break occurs within the block. If the parenthesis is not followed  | 
|
358  | 
by digits, the indentation defaults to 0. A block specified via  | 
|
359  | 
\verb|(00| is unbreakable.  | 
|
360  | 
||
361  | 
\item \verb|)| closes a pretty printing block.  | 
|
362  | 
||
363  | 
\item \verb|//| forces a line break.  | 
|
364  | 
||
365  | 
  \item \verb|/|\isa{s} allows a line break.  Here \isa{s}
 | 
|
366  | 
stands for the string of spaces (zero or more) right after the  | 
|
367  | 
  slash.  These spaces are printed if the break is \emph{not} taken.
 | 
|
368  | 
||
369  | 
  \end{description}
 | 
|
370  | 
||
371  | 
For example, the template \verb|(_ +/ _)| specifies an infix  | 
|
372  | 
operator. There are two argument positions; the delimiter  | 
|
373  | 
\verb|+| is preceded by a space and followed by a space or  | 
|
374  | 
line break; the entire phrase is a pretty printing block.  | 
|
375  | 
||
376  | 
The general idea of pretty printing with blocks and breaks is also  | 
|
377  | 
  described in \cite{paulson-ml2}.%
 | 
|
378  | 
\end{isamarkuptext}%
 | 
|
379  | 
\isamarkuptrue%  | 
|
380  | 
%  | 
|
| 35414 | 381  | 
\isamarkupsection{Explicit notation%
 | 
| 28762 | 382  | 
}  | 
383  | 
\isamarkuptrue%  | 
|
384  | 
%  | 
|
385  | 
\begin{isamarkuptext}%
 | 
|
386  | 
\begin{matharray}{rcll}
 | 
|
| 35414 | 387  | 
    \indexdef{}{command}{type\_notation}\hypertarget{command.type-notation}{\hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isacharunderscore}notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
 | 
388  | 
    \indexdef{}{command}{no\_type\_notation}\hypertarget{command.no-type-notation}{\hyperlink{command.no-type-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}type{\isacharunderscore}notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
 | 
|
| 28762 | 389  | 
    \indexdef{}{command}{notation}\hypertarget{command.notation}{\hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
 | 
390  | 
    \indexdef{}{command}{no\_notation}\hypertarget{command.no-notation}{\hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}notation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
 | 
|
| 
36508
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
391  | 
    \indexdef{}{command}{write}\hypertarget{command.write}{\hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}}} & : & \isa{{\isachardoublequote}proof{\isacharparenleft}state{\isacharparenright}\ {\isasymrightarrow}\ proof{\isacharparenleft}state{\isacharparenright}{\isachardoublequote}} \\
 | 
| 28762 | 392  | 
  \end{matharray}
 | 
393  | 
||
394  | 
  \begin{rail}
 | 
|
| 35414 | 395  | 
    ('type\_notation' | 'no\_type\_notation') target? mode? \\ (nameref mixfix + 'and')
 | 
396  | 
;  | 
|
| 29746 | 397  | 
    ('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and')
 | 
| 28762 | 398  | 
;  | 
| 
36508
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
399  | 
'write' mode? (nameref structmixfix + 'and')  | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
400  | 
;  | 
| 28762 | 401  | 
  \end{rail}
 | 
402  | 
||
403  | 
  \begin{description}
 | 
|
404  | 
||
| 35414 | 405  | 
  \item \hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isacharunderscore}notation}}}}~\isa{{\isachardoublequote}c\ {\isacharparenleft}mx{\isacharparenright}{\isachardoublequote}} associates mixfix
 | 
406  | 
syntax with an existing type constructor. The arity of the  | 
|
407  | 
constructor is retrieved from the context.  | 
|
408  | 
||
409  | 
  \item \hyperlink{command.no-type-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}type{\isacharunderscore}notation}}}} is similar to \hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isacharunderscore}notation}}}}, but removes the specified syntax annotation from
 | 
|
410  | 
the present context.  | 
|
411  | 
||
| 28762 | 412  | 
  \item \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}~\isa{{\isachardoublequote}c\ {\isacharparenleft}mx{\isacharparenright}{\isachardoublequote}} associates mixfix
 | 
| 35414 | 413  | 
syntax with an existing constant or fixed variable. The type  | 
414  | 
declaration of the given entity is retrieved from the context.  | 
|
| 28762 | 415  | 
|
416  | 
  \item \hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isacharunderscore}notation}}}} is similar to \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}},
 | 
|
417  | 
but removes the specified syntax annotation from the present  | 
|
418  | 
context.  | 
|
419  | 
||
| 
36508
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
420  | 
  \item \hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}} is similar to \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}, but
 | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
421  | 
works within an Isar proof body.  | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
422  | 
|
| 35414 | 423  | 
  \end{description}
 | 
424  | 
||
| 
36508
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
425  | 
  Note that the more primitive commands \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} and
 | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
426  | 
  \hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isacharunderscore}syntax}}}} (\secref{sec:syn-trans}) provide raw access
 | 
| 
 
03d2a2d0ee4a
allow concrete syntax for local entities within a proof body, either via regular mixfix annotations to 'fix' etc. or the separate 'write' command;
 
wenzelm 
parents: 
35414 
diff
changeset
 | 
427  | 
to the syntax tables of a global theory.%  | 
| 28762 | 428  | 
\end{isamarkuptext}%
 | 
429  | 
\isamarkuptrue%  | 
|
430  | 
%  | 
|
431  | 
\isamarkupsection{The Pure syntax \label{sec:pure-syntax}%
 | 
|
432  | 
}  | 
|
433  | 
\isamarkuptrue%  | 
|
434  | 
%  | 
|
435  | 
\isamarkupsubsection{Priority grammars \label{sec:priority-grammar}%
 | 
|
436  | 
}  | 
|
437  | 
\isamarkuptrue%  | 
|
438  | 
%  | 
|
439  | 
\begin{isamarkuptext}%
 | 
|
440  | 
A context-free grammar consists of a set of \emph{terminal
 | 
|
441  | 
  symbols}, a set of \emph{nonterminal symbols} and a set of
 | 
|
442  | 
  \emph{productions}.  Productions have the form \isa{{\isachardoublequote}A\ {\isacharequal}\ {\isasymgamma}{\isachardoublequote}},
 | 
|
443  | 
  where \isa{A} is a nonterminal and \isa{{\isasymgamma}} is a string of
 | 
|
444  | 
terminals and nonterminals. One designated nonterminal is called  | 
|
445  | 
  the \emph{root symbol}.  The language defined by the grammar
 | 
|
446  | 
consists of all strings of terminals that can be derived from the  | 
|
447  | 
root symbol by applying productions as rewrite rules.  | 
|
448  | 
||
449  | 
  The standard Isabelle parser for inner syntax uses a \emph{priority
 | 
|
450  | 
grammar}. Each nonterminal is decorated by an integer priority:  | 
|
451  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}{\isachardoublequote}}.  In a derivation, \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}{\isachardoublequote}} may be rewritten
 | 
|
452  | 
  using a production \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup q\isactrlsup {\isacharparenright}\ {\isacharequal}\ {\isasymgamma}{\isachardoublequote}} only if \isa{{\isachardoublequote}p\ {\isasymle}\ q{\isachardoublequote}}.  Any
 | 
|
453  | 
priority grammar can be translated into a normal context-free  | 
|
454  | 
grammar by introducing new nonterminals and productions.  | 
|
455  | 
||
456  | 
  \medskip Formally, a set of context free productions \isa{G}
 | 
|
457  | 
  induces a derivation relation \isa{{\isachardoublequote}{\isasymlongrightarrow}\isactrlsub G{\isachardoublequote}} as follows.  Let \isa{{\isasymalpha}} and \isa{{\isasymbeta}} denote strings of terminal or nonterminal symbols.
 | 
|
458  | 
  Then \isa{{\isachardoublequote}{\isasymalpha}\ A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}\ {\isasymbeta}\ {\isasymlongrightarrow}\isactrlsub G\ {\isasymalpha}\ {\isasymgamma}\ {\isasymbeta}{\isachardoublequote}} holds if and only if \isa{G}
 | 
|
459  | 
  contains some production \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup q\isactrlsup {\isacharparenright}\ {\isacharequal}\ {\isasymgamma}{\isachardoublequote}} for \isa{{\isachardoublequote}p\ {\isasymle}\ q{\isachardoublequote}}.
 | 
|
460  | 
||
461  | 
\medskip The following grammar for arithmetic expressions  | 
|
462  | 
demonstrates how binding power and associativity of operators can be  | 
|
463  | 
enforced by priorities.  | 
|
464  | 
||
465  | 
  \begin{center}
 | 
|
466  | 
  \begin{tabular}{rclr}
 | 
|
467  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \verb|(| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|)| \\
 | 
|
468  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \verb|0| \\
 | 
|
469  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|+| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \\
 | 
|
470  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|*| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \\
 | 
|
471  | 
  \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \verb|-| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \\
 | 
|
472  | 
  \end{tabular}
 | 
|
473  | 
  \end{center}
 | 
|
474  | 
The choice of priorities determines that \verb|-| binds  | 
|
475  | 
tighter than \verb|*|, which binds tighter than \verb|+|. Furthermore \verb|+| associates to the left and  | 
|
476  | 
\verb|*| to the right.  | 
|
477  | 
||
478  | 
\medskip For clarity, grammars obey these conventions:  | 
|
479  | 
  \begin{itemize}
 | 
|
480  | 
||
481  | 
\item All priorities must lie between 0 and 1000.  | 
|
482  | 
||
483  | 
\item Priority 0 on the right-hand side and priority 1000 on the  | 
|
484  | 
left-hand side may be omitted.  | 
|
485  | 
||
486  | 
  \item The production \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}\ {\isacharequal}\ {\isasymalpha}{\isachardoublequote}} is written as \isa{{\isachardoublequote}A\ {\isacharequal}\ {\isasymalpha}\ {\isacharparenleft}p{\isacharparenright}{\isachardoublequote}}, i.e.\ the priority of the left-hand side actually appears in
 | 
|
487  | 
a column on the far right.  | 
|
488  | 
||
489  | 
  \item Alternatives are separated by \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}}.
 | 
|
490  | 
||
491  | 
  \item Repetition is indicated by dots \isa{{\isachardoublequote}{\isacharparenleft}{\isasymdots}{\isacharparenright}{\isachardoublequote}} in an informal
 | 
|
492  | 
but obvious way.  | 
|
493  | 
||
494  | 
  \end{itemize}
 | 
|
495  | 
||
496  | 
Using these conventions, the example grammar specification above  | 
|
497  | 
takes the form:  | 
|
498  | 
  \begin{center}
 | 
|
499  | 
  \begin{tabular}{rclc}
 | 
|
500  | 
    \isa{A} & \isa{{\isachardoublequote}{\isacharequal}{\isachardoublequote}} & \verb|(| \isa{A} \verb|)| \\
 | 
|
501  | 
              & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|0| & \qquad\qquad \\
 | 
|
502  | 
              & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{A} \verb|+| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
503  | 
              & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|*| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
504  | 
              & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|-| \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
505  | 
  \end{tabular}
 | 
|
506  | 
  \end{center}%
 | 
|
507  | 
\end{isamarkuptext}%
 | 
|
508  | 
\isamarkuptrue%  | 
|
509  | 
%  | 
|
510  | 
\isamarkupsubsection{The Pure grammar%
 | 
|
511  | 
}  | 
|
512  | 
\isamarkuptrue%  | 
|
513  | 
%  | 
|
514  | 
\begin{isamarkuptext}%
 | 
|
515  | 
The priority grammar of the \isa{{\isachardoublequote}Pure{\isachardoublequote}} theory is defined as follows:
 | 
|
516  | 
||
517  | 
%FIXME syntax for "index" (?)  | 
|
518  | 
%FIXME "op" versions of ==> etc. (?)  | 
|
519  | 
||
520  | 
  \begin{center}
 | 
|
521  | 
  \begin{supertabular}{rclr}
 | 
|
522  | 
||
523  | 
  \indexdef{inner}{syntax}{any}\hypertarget{syntax.inner.any}{\hyperlink{syntax.inner.any}{\mbox{\isa{any}}}} & = & \isa{{\isachardoublequote}prop\ \ {\isacharbar}\ \ logic{\isachardoublequote}} \\\\
 | 
|
524  | 
||
525  | 
  \indexdef{inner}{syntax}{prop}\hypertarget{syntax.inner.prop}{\hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}} & = & \verb|(| \isa{prop} \verb|)| \\
 | 
|
526  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{4}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|::| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
527  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|=?=| \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
528  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|==| \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
529  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \isa{{\isachardoublequote}{\isasymequiv}{\isachardoublequote}} \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
| 28857 | 530  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|&&&| \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{2}}{\isacharparenright}{\isachardoublequote}} \\
 | 
| 28762 | 531  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|==>| \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}} \\
 | 
532  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{2}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}} \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
533  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|[|\verb,|,\verb|| \isa{prop} \verb|;| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|;| \isa{prop} \verb||\verb,|,\verb|]| \verb|==>| \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
534  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymlbrakk}{\isachardoublequote}} \isa{prop} \verb|;| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|;| \isa{prop} \isa{{\isachardoublequote}{\isasymrbrakk}{\isachardoublequote}} \isa{{\isachardoublequote}{\isasymLongrightarrow}{\isachardoublequote}} \isa{{\isachardoublequote}prop\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
535  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|!!| \isa{idts} \verb|.| \isa{prop} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
536  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}} \isa{idts} \verb|.| \isa{prop} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
537  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|OFCLASS| \verb|(| \isa{type} \verb|,| \isa{logic} \verb|)| \\
 | 
|
538  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|SORT_CONSTRAINT| \verb|(| \isa{type} \verb|)| \\
 | 
|
| 28857 | 539  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|TERM| \isa{logic} \\
 | 
| 28762 | 540  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|PROP| \isa{aprop} \\\\
 | 
541  | 
||
| 28857 | 542  | 
  \indexdef{inner}{syntax}{aprop}\hypertarget{syntax.inner.aprop}{\hyperlink{syntax.inner.aprop}{\mbox{\isa{aprop}}}} & = & \verb|(| \isa{aprop} \verb|)| \\
 | 
543  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ longid\ \ {\isacharbar}\ \ var\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_|\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|...| \\
 | 
|
544  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|CONST| \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|CONST| \isa{{\isachardoublequote}longid{\isachardoublequote}} \\
 | 
|
| 28762 | 545  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}logic\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ \ any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ {\isasymdots}\ any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
546  | 
||
547  | 
  \indexdef{inner}{syntax}{logic}\hypertarget{syntax.inner.logic}{\hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}}} & = & \verb|(| \isa{logic} \verb|)| \\
 | 
|
548  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}logic\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{4}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|::| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
549  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ longid\ \ {\isacharbar}\ \ var\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_|\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|...| \\
 | 
|
| 28857 | 550  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|CONST| \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|CONST| \isa{{\isachardoublequote}longid{\isachardoublequote}} \\
 | 
| 28762 | 551  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}logic\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ \ any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ {\isasymdots}\ any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}{\isachardoublequote}} \\
 | 
552  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|%| \isa{pttrns} \verb|.| \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
553  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isasymlambda}} \isa{pttrns} \verb|.| \isa{{\isachardoublequote}any\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{3}}\isactrlsup {\isacharparenright}{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
554  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|TYPE| \verb|(| \isa{type} \verb|)| \\\\
 | 
|
555  | 
||
556  | 
  \indexdef{inner}{syntax}{idt}\hypertarget{syntax.inner.idt}{\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}}} & = & \verb|(| \isa{idt} \verb|)|\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ id\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_| \\
 | 
|
557  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{id} \verb|::| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
558  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|_| \verb|::| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
|
559  | 
||
560  | 
  \indexdef{inner}{syntax}{idts}\hypertarget{syntax.inner.idts}{\hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}} & = & \isa{{\isachardoublequote}idt\ \ {\isacharbar}\ \ idt\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}\ idts{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
|
561  | 
||
562  | 
  \indexdef{inner}{syntax}{pttrn}\hypertarget{syntax.inner.pttrn}{\hyperlink{syntax.inner.pttrn}{\mbox{\isa{pttrn}}}} & = & \isa{idt} \\\\
 | 
|
563  | 
||
564  | 
  \indexdef{inner}{syntax}{pttrns}\hypertarget{syntax.inner.pttrns}{\hyperlink{syntax.inner.pttrns}{\mbox{\isa{pttrns}}}} & = & \isa{{\isachardoublequote}pttrn\ \ {\isacharbar}\ \ pttrn\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}\ pttrns{\isachardoublequote}} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
|
565  | 
||
566  | 
  \indexdef{inner}{syntax}{type}\hypertarget{syntax.inner.type}{\hyperlink{syntax.inner.type}{\mbox{\isa{type}}}} & = & \verb|(| \isa{type} \verb|)| \\
 | 
|
567  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}tid\ \ {\isacharbar}\ \ tvar\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_| \\
 | 
|
568  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}tid{\isachardoublequote}} \verb|::| \isa{{\isachardoublequote}sort\ \ {\isacharbar}\ \ tvar\ \ {\isachardoublequote}}\verb|::| \isa{{\isachardoublequote}sort\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|_| \verb|::| \isa{{\isachardoublequote}sort{\isachardoublequote}} \\
 | 
|
569  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ type\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ id\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|(| \isa{type} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{type} \verb|)| \isa{id} \\
 | 
|
| 29746 | 570  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}longid\ \ {\isacharbar}\ \ type\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isadigit{0}}\isactrlsup {\isacharparenright}\ longid{\isachardoublequote}} \\
 | 
571  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|(| \isa{type} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{type} \verb|)| \isa{longid} \\
 | 
|
| 28762 | 572  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}type\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \verb|=>| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
573  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \isa{{\isachardoublequote}type\isactrlsup {\isacharparenleft}\isactrlsup {\isadigit{1}}\isactrlsup {\isacharparenright}{\isachardoublequote}} \isa{{\isachardoublequote}{\isasymRightarrow}{\isachardoublequote}} \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
574  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|[| \isa{type} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{type} \verb|]| \verb|=>| \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\
 | 
|
575  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|[| \isa{type} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{type} \verb|]| \isa{{\isachardoublequote}{\isasymRightarrow}{\isachardoublequote}} \isa{type} & \isa{{\isachardoublequote}{\isacharparenleft}{\isadigit{0}}{\isacharparenright}{\isachardoublequote}} \\\\
 | 
|
576  | 
||
| 29746 | 577  | 
  \indexdef{inner}{syntax}{sort}\hypertarget{syntax.inner.sort}{\hyperlink{syntax.inner.sort}{\mbox{\isa{sort}}}} & = & \isa{{\isachardoublequote}id\ \ {\isacharbar}\ \ longid\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|{}| \\
 | 
578  | 
    & \isa{{\isachardoublequote}{\isacharbar}{\isachardoublequote}} & \verb|{| \isa{{\isachardoublequote}{\isacharparenleft}id\ {\isacharbar}\ longid{\isacharparenright}{\isachardoublequote}} \verb|,| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|,| \isa{{\isachardoublequote}{\isacharparenleft}id\ {\isacharbar}\ longid{\isacharparenright}{\isachardoublequote}} \verb|}| \\
 | 
|
| 28762 | 579  | 
  \end{supertabular}
 | 
580  | 
  \end{center}
 | 
|
581  | 
||
582  | 
\medskip Here literal terminals are printed \verb|verbatim|;  | 
|
583  | 
  see also \secref{sec:inner-lex} for further token categories of the
 | 
|
584  | 
inner syntax. The meaning of the nonterminals defined by the above  | 
|
585  | 
grammar is as follows:  | 
|
586  | 
||
587  | 
  \begin{description}
 | 
|
588  | 
||
589  | 
  \item \indexref{inner}{syntax}{any}\hyperlink{syntax.inner.any}{\mbox{\isa{any}}} denotes any term.
 | 
|
590  | 
||
591  | 
  \item \indexref{inner}{syntax}{prop}\hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}} denotes meta-level propositions,
 | 
|
592  | 
  which are terms of type \isa{prop}.  The syntax of such formulae of
 | 
|
593  | 
the meta-logic is carefully distinguished from usual conventions for  | 
|
594  | 
  object-logics.  In particular, plain \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}}-term notation is
 | 
|
595  | 
  \emph{not} recognized as \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}.
 | 
|
596  | 
||
597  | 
  \item \indexref{inner}{syntax}{aprop}\hyperlink{syntax.inner.aprop}{\mbox{\isa{aprop}}} denotes atomic propositions, which
 | 
|
598  | 
  are embedded into regular \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}} by means of an
 | 
|
599  | 
explicit \verb|PROP| token.  | 
|
600  | 
||
601  | 
  Terms of type \isa{prop} with non-constant head, e.g.\ a plain
 | 
|
602  | 
  variable, are printed in this form.  Constants that yield type \isa{prop} are expected to provide their own concrete syntax; otherwise
 | 
|
603  | 
  the printed version will appear like \hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}} and
 | 
|
604  | 
  cannot be parsed again as \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}.
 | 
|
605  | 
||
606  | 
  \item \indexref{inner}{syntax}{logic}\hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}} denotes arbitrary terms of a
 | 
|
607  | 
  logical type, excluding type \isa{prop}.  This is the main
 | 
|
608  | 
  syntactic category of object-logic entities, covering plain \isa{{\isasymlambda}}-term notation (variables, abstraction, application), plus
 | 
|
609  | 
anything defined by the user.  | 
|
610  | 
||
611  | 
When specifying notation for logical entities, all logical types  | 
|
612  | 
  (excluding \isa{prop}) are \emph{collapsed} to this single category
 | 
|
613  | 
  of \hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}}.
 | 
|
614  | 
||
615  | 
  \item \indexref{inner}{syntax}{idt}\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}} denotes identifiers, possibly
 | 
|
616  | 
constrained by types.  | 
|
617  | 
||
618  | 
  \item \indexref{inner}{syntax}{idts}\hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}} denotes a sequence of \indexref{inner}{syntax}{idt}\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}}.  This is the most basic category for variables in
 | 
|
619  | 
  iterated binders, such as \isa{{\isachardoublequote}{\isasymlambda}{\isachardoublequote}} or \isa{{\isachardoublequote}{\isasymAnd}{\isachardoublequote}}.
 | 
|
620  | 
||
621  | 
  \item \indexref{inner}{syntax}{pttrn}\hyperlink{syntax.inner.pttrn}{\mbox{\isa{pttrn}}} and \indexref{inner}{syntax}{pttrns}\hyperlink{syntax.inner.pttrns}{\mbox{\isa{pttrns}}}
 | 
|
622  | 
denote patterns for abstraction, cases bindings etc. In Pure, these  | 
|
623  | 
  categories start as a merely copy of \hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}} and
 | 
|
624  | 
  \hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}, respectively.  Object-logics may add
 | 
|
625  | 
additional productions for binding forms.  | 
|
626  | 
||
627  | 
  \item \indexref{inner}{syntax}{type}\hyperlink{syntax.inner.type}{\mbox{\isa{type}}} denotes types of the meta-logic.
 | 
|
628  | 
||
629  | 
  \item \indexref{inner}{syntax}{sort}\hyperlink{syntax.inner.sort}{\mbox{\isa{sort}}} denotes meta-level sorts.
 | 
|
630  | 
||
631  | 
  \end{description}
 | 
|
632  | 
||
633  | 
Here are some further explanations of certain syntax features.  | 
|
634  | 
||
635  | 
  \begin{itemize}
 | 
|
636  | 
||
637  | 
  \item In \hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}, note that \isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ nat\ y{\isachardoublequote}} is
 | 
|
638  | 
  parsed as \isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}nat\ y{\isacharparenright}{\isachardoublequote}}, treating \isa{y} like a type
 | 
|
639  | 
  constructor applied to \isa{nat}.  To avoid this interpretation,
 | 
|
640  | 
  write \isa{{\isachardoublequote}{\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}\ y{\isachardoublequote}} with explicit parentheses.
 | 
|
641  | 
||
642  | 
  \item Similarly, \isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ nat\ y\ {\isacharcolon}{\isacharcolon}\ nat{\isachardoublequote}} is parsed as \isa{{\isachardoublequote}x\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}nat\ y\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}{\isachardoublequote}}.  The correct form is \isa{{\isachardoublequote}{\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}\ {\isacharparenleft}y\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}{\isachardoublequote}}, or \isa{{\isachardoublequote}{\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}\ y\ {\isacharcolon}{\isacharcolon}\ nat{\isachardoublequote}} if \isa{y} is last in the
 | 
|
643  | 
sequence of identifiers.  | 
|
644  | 
||
645  | 
\item Type constraints for terms bind very weakly. For example,  | 
|
646  | 
  \isa{{\isachardoublequote}x\ {\isacharless}\ y\ {\isacharcolon}{\isacharcolon}\ nat{\isachardoublequote}} is normally parsed as \isa{{\isachardoublequote}{\isacharparenleft}x\ {\isacharless}\ y{\isacharparenright}\ {\isacharcolon}{\isacharcolon}\ nat{\isachardoublequote}}, unless \isa{{\isachardoublequote}{\isacharless}{\isachardoublequote}} has a very low priority, in which case the
 | 
|
647  | 
  input is likely to be ambiguous.  The correct form is \isa{{\isachardoublequote}x\ {\isacharless}\ {\isacharparenleft}y\ {\isacharcolon}{\isacharcolon}\ nat{\isacharparenright}{\isachardoublequote}}.
 | 
|
648  | 
||
649  | 
\item Constraints may be either written with two literal colons  | 
|
650  | 
``\verb|::|'' or the double-colon symbol \verb|\<Colon>|,  | 
|
651  | 
  which actually looks exactly the same in some {\LaTeX} styles.
 | 
|
652  | 
||
653  | 
\item Dummy variables (written as underscore) may occur in different  | 
|
654  | 
roles.  | 
|
655  | 
||
656  | 
  \begin{description}
 | 
|
657  | 
||
658  | 
  \item A type ``\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}'' or ``\isa{{\isachardoublequote}{\isacharunderscore}\ {\isacharcolon}{\isacharcolon}\ sort{\isachardoublequote}}'' acts like an
 | 
|
659  | 
anonymous inference parameter, which is filled-in according to the  | 
|
660  | 
most general type produced by the type-checking phase.  | 
|
661  | 
||
662  | 
  \item A bound ``\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}'' refers to a vacuous abstraction, where
 | 
|
663  | 
the body does not refer to the binding introduced here. As in the  | 
|
664  | 
  term \isa{{\isachardoublequote}{\isasymlambda}x\ {\isacharunderscore}{\isachardot}\ x{\isachardoublequote}}, which is \isa{{\isachardoublequote}{\isasymalpha}{\isachardoublequote}}-equivalent to \isa{{\isachardoublequote}{\isasymlambda}x\ y{\isachardot}\ x{\isachardoublequote}}.
 | 
|
665  | 
||
666  | 
  \item A free ``\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}'' refers to an implicit outer binding.
 | 
|
667  | 
  Higher definitional packages usually allow forms like \isa{{\isachardoublequote}f\ x\ {\isacharunderscore}\ {\isacharequal}\ x{\isachardoublequote}}.
 | 
|
668  | 
||
669  | 
  \item A schematic ``\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}'' (within a term pattern, see
 | 
|
670  | 
  \secref{sec:term-decls}) refers to an anonymous variable that is
 | 
|
671  | 
implicitly abstracted over its context of locally bound variables.  | 
|
672  | 
  For example, this allows pattern matching of \isa{{\isachardoublequote}{\isacharbraceleft}x{\isachardot}\ f\ x\ {\isacharequal}\ g\ x{\isacharbraceright}{\isachardoublequote}} against \isa{{\isachardoublequote}{\isacharbraceleft}x{\isachardot}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharunderscore}{\isacharbraceright}{\isachardoublequote}}, or even \isa{{\isachardoublequote}{\isacharbraceleft}{\isacharunderscore}{\isachardot}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharunderscore}{\isacharbraceright}{\isachardoublequote}} by
 | 
|
673  | 
using both bound and schematic dummies.  | 
|
674  | 
||
675  | 
  \end{description}
 | 
|
676  | 
||
677  | 
\item The three literal dots ``\verb|...|'' may be also  | 
|
678  | 
written as ellipsis symbol \verb|\<dots>|. In both cases this  | 
|
679  | 
refers to a special schematic variable, which is bound in the  | 
|
680  | 
context. This special term abbreviation works nicely with  | 
|
681  | 
  calculational reasoning (\secref{sec:calculation}).
 | 
|
682  | 
||
683  | 
  \end{itemize}%
 | 
|
684  | 
\end{isamarkuptext}%
 | 
|
685  | 
\isamarkuptrue%  | 
|
686  | 
%  | 
|
687  | 
\isamarkupsection{Lexical matters \label{sec:inner-lex}%
 | 
|
688  | 
}  | 
|
689  | 
\isamarkuptrue%  | 
|
690  | 
%  | 
|
691  | 
\begin{isamarkuptext}%
 | 
|
692  | 
The inner lexical syntax vaguely resembles the outer one  | 
|
693  | 
  (\secref{sec:outer-lex}), but some details are different.  There are
 | 
|
694  | 
two main categories of inner syntax tokens:  | 
|
695  | 
||
696  | 
  \begin{enumerate}
 | 
|
697  | 
||
698  | 
  \item \emph{delimiters} --- the literal tokens occurring in
 | 
|
699  | 
productions of the given priority grammar (cf.\  | 
|
700  | 
  \secref{sec:priority-grammar});
 | 
|
701  | 
||
702  | 
  \item \emph{named tokens} --- various categories of identifiers etc.
 | 
|
703  | 
||
704  | 
  \end{enumerate}
 | 
|
705  | 
||
706  | 
Delimiters override named tokens and may thus render certain  | 
|
707  | 
identifiers inaccessible. Sometimes the logical context admits  | 
|
708  | 
alternative ways to refer to the same entity, potentially via  | 
|
709  | 
qualified names.  | 
|
710  | 
||
711  | 
\medskip The categories for named tokens are defined once and for  | 
|
712  | 
all as follows, reusing some categories of the outer token syntax  | 
|
713  | 
  (\secref{sec:outer-lex}).
 | 
|
714  | 
||
715  | 
  \begin{center}
 | 
|
716  | 
  \begin{supertabular}{rcl}
 | 
|
717  | 
    \indexdef{inner}{syntax}{id}\hypertarget{syntax.inner.id}{\hyperlink{syntax.inner.id}{\mbox{\isa{id}}}} & = & \indexref{}{syntax}{ident}\hyperlink{syntax.ident}{\mbox{\isa{ident}}} \\
 | 
|
718  | 
    \indexdef{inner}{syntax}{longid}\hypertarget{syntax.inner.longid}{\hyperlink{syntax.inner.longid}{\mbox{\isa{longid}}}} & = & \indexref{}{syntax}{longident}\hyperlink{syntax.longident}{\mbox{\isa{longident}}} \\
 | 
|
719  | 
    \indexdef{inner}{syntax}{var}\hypertarget{syntax.inner.var}{\hyperlink{syntax.inner.var}{\mbox{\isa{var}}}} & = & \indexref{}{syntax}{var}\hyperlink{syntax.var}{\mbox{\isa{var}}} \\
 | 
|
720  | 
    \indexdef{inner}{syntax}{tid}\hypertarget{syntax.inner.tid}{\hyperlink{syntax.inner.tid}{\mbox{\isa{tid}}}} & = & \indexref{}{syntax}{typefree}\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}} \\
 | 
|
721  | 
    \indexdef{inner}{syntax}{tvar}\hypertarget{syntax.inner.tvar}{\hyperlink{syntax.inner.tvar}{\mbox{\isa{tvar}}}} & = & \indexref{}{syntax}{typevar}\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}} \\
 | 
|
722  | 
    \indexdef{inner}{syntax}{num}\hypertarget{syntax.inner.num}{\hyperlink{syntax.inner.num}{\mbox{\isa{num}}}} & = & \indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\
 | 
|
| 29158 | 723  | 
    \indexdef{inner}{syntax}{float\_token}\hypertarget{syntax.inner.float-token}{\hyperlink{syntax.inner.float-token}{\mbox{\isa{float{\isacharunderscore}token}}}} & = & \indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\
 | 
| 28762 | 724  | 
    \indexdef{inner}{syntax}{xnum}\hypertarget{syntax.inner.xnum}{\hyperlink{syntax.inner.xnum}{\mbox{\isa{xnum}}}} & = & \verb|#|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|#-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\
 | 
725  | 
||
726  | 
    \indexdef{inner}{syntax}{xstr}\hypertarget{syntax.inner.xstr}{\hyperlink{syntax.inner.xstr}{\mbox{\isa{xstr}}}} & = & \verb|''| \isa{{\isachardoublequote}{\isasymdots}{\isachardoublequote}} \verb|''| \\
 | 
|
727  | 
  \end{supertabular}
 | 
|
728  | 
  \end{center}
 | 
|
729  | 
||
| 29158 | 730  | 
  The token categories \hyperlink{syntax.inner.num}{\mbox{\isa{num}}}, \hyperlink{syntax.inner.float-token}{\mbox{\isa{float{\isacharunderscore}token}}}, \hyperlink{syntax.inner.xnum}{\mbox{\isa{xnum}}}, and \hyperlink{syntax.inner.xstr}{\mbox{\isa{xstr}}} are
 | 
731  | 
not used in Pure. Object-logics may implement numerals and string  | 
|
732  | 
constants by adding appropriate syntax declarations, together with  | 
|
733  | 
some translation functions (e.g.\ see Isabelle/HOL).  | 
|
734  | 
||
735  | 
  The derived categories \indexdef{inner}{syntax}{num\_const}\hypertarget{syntax.inner.num-const}{\hyperlink{syntax.inner.num-const}{\mbox{\isa{num{\isacharunderscore}const}}}} and
 | 
|
736  | 
  \indexdef{inner}{syntax}{float\_const}\hypertarget{syntax.inner.float-const}{\hyperlink{syntax.inner.float-const}{\mbox{\isa{float{\isacharunderscore}const}}}} provide robust access to \hyperlink{syntax.inner.num}{\mbox{\isa{num}}}, and \hyperlink{syntax.inner.float-token}{\mbox{\isa{float{\isacharunderscore}token}}}, respectively: the
 | 
|
737  | 
syntax tree holds a syntactic constant instead of a free variable.%  | 
|
| 28762 | 738  | 
\end{isamarkuptext}%
 | 
739  | 
\isamarkuptrue%  | 
|
740  | 
%  | 
|
741  | 
\isamarkupsection{Syntax and translations \label{sec:syn-trans}%
 | 
|
742  | 
}  | 
|
743  | 
\isamarkuptrue%  | 
|
744  | 
%  | 
|
745  | 
\begin{isamarkuptext}%
 | 
|
746  | 
\begin{matharray}{rcl}
 | 
|
747  | 
    \indexdef{}{command}{nonterminals}\hypertarget{command.nonterminals}{\hyperlink{command.nonterminals}{\mbox{\isa{\isacommand{nonterminals}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
748  | 
    \indexdef{}{command}{syntax}\hypertarget{command.syntax}{\hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
749  | 
    \indexdef{}{command}{no\_syntax}\hypertarget{command.no-syntax}{\hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isacharunderscore}syntax}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
750  | 
    \indexdef{}{command}{translations}\hypertarget{command.translations}{\hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
751  | 
    \indexdef{}{command}{no\_translations}\hypertarget{command.no-translations}{\hyperlink{command.no-translations}{\mbox{\isa{\isacommand{no{\isacharunderscore}translations}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
752  | 
  \end{matharray}
 | 
|
753  | 
||
754  | 
  \begin{rail}
 | 
|
755  | 
'nonterminals' (name +)  | 
|
756  | 
;  | 
|
757  | 
    ('syntax' | 'no\_syntax') mode? (constdecl +)
 | 
|
758  | 
;  | 
|
759  | 
    ('translations' | 'no\_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
 | 
|
760  | 
;  | 
|
761  | 
||
762  | 
    mode: ('(' ( name | 'output' | name 'output' ) ')')
 | 
|
763  | 
;  | 
|
764  | 
    transpat: ('(' nameref ')')? string
 | 
|
765  | 
;  | 
|
766  | 
  \end{rail}
 | 
|
767  | 
||
768  | 
  \begin{description}
 | 
|
769  | 
||
770  | 
  \item \hyperlink{command.nonterminals}{\mbox{\isa{\isacommand{nonterminals}}}}~\isa{c} declares a type
 | 
|
771  | 
  constructor \isa{c} (without arguments) to act as purely syntactic
 | 
|
772  | 
type: a nonterminal symbol of the inner syntax.  | 
|
773  | 
||
774  | 
  \item \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}~\isa{{\isachardoublequote}{\isacharparenleft}mode{\isacharparenright}\ decls{\isachardoublequote}} is similar to
 | 
|
775  | 
  \hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}~\isa{decls}, except that the actual logical
 | 
|
776  | 
signature extension is omitted. Thus the context free grammar of  | 
|
777  | 
Isabelle's inner syntax may be augmented in arbitrary ways,  | 
|
778  | 
  independently of the logic.  The \isa{mode} argument refers to the
 | 
|
779  | 
  print mode that the grammar rules belong; unless the \indexref{}{keyword}{output}\hyperlink{keyword.output}{\mbox{\isa{\isakeyword{output}}}} indicator is given, all productions are added both to the
 | 
|
780  | 
input and output grammar.  | 
|
781  | 
||
782  | 
  \item \hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isacharunderscore}syntax}}}}~\isa{{\isachardoublequote}{\isacharparenleft}mode{\isacharparenright}\ decls{\isachardoublequote}} removes grammar
 | 
|
783  | 
  declarations (and translations) resulting from \isa{decls}, which
 | 
|
784  | 
  are interpreted in the same manner as for \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} above.
 | 
|
785  | 
||
786  | 
  \item \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}~\isa{rules} specifies syntactic
 | 
|
787  | 
  translation rules (i.e.\ macros): parse~/ print rules (\isa{{\isachardoublequote}{\isasymrightleftharpoons}{\isachardoublequote}}),
 | 
|
788  | 
  parse rules (\isa{{\isachardoublequote}{\isasymrightharpoonup}{\isachardoublequote}}), or print rules (\isa{{\isachardoublequote}{\isasymleftharpoondown}{\isachardoublequote}}).
 | 
|
789  | 
Translation patterns may be prefixed by the syntactic category to be  | 
|
790  | 
  used for parsing; the default is \isa{logic}.
 | 
|
791  | 
||
792  | 
  \item \hyperlink{command.no-translations}{\mbox{\isa{\isacommand{no{\isacharunderscore}translations}}}}~\isa{rules} removes syntactic
 | 
|
793  | 
translation rules, which are interpreted in the same manner as for  | 
|
794  | 
  \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}} above.
 | 
|
795  | 
||
796  | 
  \end{description}%
 | 
|
797  | 
\end{isamarkuptext}%
 | 
|
798  | 
\isamarkuptrue%  | 
|
799  | 
%  | 
|
800  | 
\isamarkupsection{Syntax translation functions \label{sec:tr-funs}%
 | 
|
801  | 
}  | 
|
802  | 
\isamarkuptrue%  | 
|
803  | 
%  | 
|
804  | 
\begin{isamarkuptext}%
 | 
|
805  | 
\begin{matharray}{rcl}
 | 
|
806  | 
    \indexdef{}{command}{parse\_ast\_translation}\hypertarget{command.parse-ast-translation}{\hyperlink{command.parse-ast-translation}{\mbox{\isa{\isacommand{parse{\isacharunderscore}ast{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
807  | 
    \indexdef{}{command}{parse\_translation}\hypertarget{command.parse-translation}{\hyperlink{command.parse-translation}{\mbox{\isa{\isacommand{parse{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
808  | 
    \indexdef{}{command}{print\_translation}\hypertarget{command.print-translation}{\hyperlink{command.print-translation}{\mbox{\isa{\isacommand{print{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
809  | 
    \indexdef{}{command}{typed\_print\_translation}\hypertarget{command.typed-print-translation}{\hyperlink{command.typed-print-translation}{\mbox{\isa{\isacommand{typed{\isacharunderscore}print{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
810  | 
    \indexdef{}{command}{print\_ast\_translation}\hypertarget{command.print-ast-translation}{\hyperlink{command.print-ast-translation}{\mbox{\isa{\isacommand{print{\isacharunderscore}ast{\isacharunderscore}translation}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
 | 
|
811  | 
  \end{matharray}
 | 
|
812  | 
||
813  | 
  \begin{rail}
 | 
|
814  | 
( 'parse\_ast\_translation' | 'parse\_translation' | 'print\_translation' |  | 
|
815  | 
    'typed\_print\_translation' | 'print\_ast\_translation' ) ('(advanced)')? text
 | 
|
816  | 
;  | 
|
817  | 
  \end{rail}
 | 
|
818  | 
||
819  | 
Syntax translation functions written in ML admit almost arbitrary  | 
|
820  | 
manipulations of Isabelle's inner syntax. Any of the above commands  | 
|
821  | 
  have a single \railqtok{text} argument that refers to an ML
 | 
|
822  | 
expression of appropriate type, which are as follows by default:  | 
|
823  | 
||
824  | 
%FIXME proper antiquotations  | 
|
825  | 
\begin{ttbox}
 | 
|
826  | 
val parse_ast_translation : (string * (ast list -> ast)) list  | 
|
827  | 
val parse_translation : (string * (term list -> term)) list  | 
|
828  | 
val print_translation : (string * (term list -> term)) list  | 
|
829  | 
val typed_print_translation :  | 
|
830  | 
(string * (bool -> typ -> term list -> term)) list  | 
|
831  | 
val print_ast_translation : (string * (ast list -> ast)) list  | 
|
832  | 
\end{ttbox}
 | 
|
833  | 
||
834  | 
  If the \isa{{\isachardoublequote}{\isacharparenleft}advanced{\isacharparenright}{\isachardoublequote}} option is given, the corresponding
 | 
|
835  | 
translation functions may depend on the current theory or proof  | 
|
836  | 
context. This allows to implement advanced syntax mechanisms, as  | 
|
837  | 
translations functions may refer to specific theory declarations or  | 
|
838  | 
auxiliary proof data.  | 
|
839  | 
||
| 30397 | 840  | 
  See also \cite{isabelle-ref} for more information on the general
 | 
841  | 
concept of syntax transformations in Isabelle.  | 
|
| 28762 | 842  | 
|
843  | 
%FIXME proper antiquotations  | 
|
844  | 
\begin{ttbox}
 | 
|
845  | 
val parse_ast_translation:  | 
|
846  | 
(string * (Proof.context -> ast list -> ast)) list  | 
|
847  | 
val parse_translation:  | 
|
848  | 
(string * (Proof.context -> term list -> term)) list  | 
|
849  | 
val print_translation:  | 
|
850  | 
(string * (Proof.context -> term list -> term)) list  | 
|
851  | 
val typed_print_translation:  | 
|
852  | 
(string * (Proof.context -> bool -> typ -> term list -> term)) list  | 
|
853  | 
val print_ast_translation:  | 
|
854  | 
(string * (Proof.context -> ast list -> ast)) list  | 
|
855  | 
\end{ttbox}%
 | 
|
856  | 
\end{isamarkuptext}%
 | 
|
857  | 
\isamarkuptrue%  | 
|
858  | 
%  | 
|
859  | 
\isamarkupsection{Inspecting the syntax%
 | 
|
860  | 
}  | 
|
861  | 
\isamarkuptrue%  | 
|
862  | 
%  | 
|
863  | 
\begin{isamarkuptext}%
 | 
|
864  | 
\begin{matharray}{rcl}
 | 
|
865  | 
    \indexdef{}{command}{print\_syntax}\hypertarget{command.print-syntax}{\hyperlink{command.print-syntax}{\mbox{\isa{\isacommand{print{\isacharunderscore}syntax}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
 | 
|
866  | 
  \end{matharray}
 | 
|
867  | 
||
868  | 
  \begin{description}
 | 
|
869  | 
||
870  | 
  \item \hyperlink{command.print-syntax}{\mbox{\isa{\isacommand{print{\isacharunderscore}syntax}}}} prints the inner syntax of the
 | 
|
871  | 
current context. The output can be quite large; the most important  | 
|
872  | 
sections are explained below.  | 
|
873  | 
||
874  | 
  \begin{description}
 | 
|
875  | 
||
876  | 
  \item \isa{{\isachardoublequote}lexicon{\isachardoublequote}} lists the delimiters of the inner token
 | 
|
877  | 
  language; see \secref{sec:inner-lex}.
 | 
|
878  | 
||
879  | 
  \item \isa{{\isachardoublequote}prods{\isachardoublequote}} lists the productions of the underlying
 | 
|
880  | 
  priority grammar; see \secref{sec:priority-grammar}.
 | 
|
881  | 
||
882  | 
  The nonterminal \isa{{\isachardoublequote}A\isactrlsup {\isacharparenleft}\isactrlsup p\isactrlsup {\isacharparenright}{\isachardoublequote}} is rendered in plain text as \isa{{\isachardoublequote}A{\isacharbrackleft}p{\isacharbrackright}{\isachardoublequote}}; delimiters are quoted.  Many productions have an extra
 | 
|
883  | 
  \isa{{\isachardoublequote}{\isasymdots}\ {\isacharequal}{\isachargreater}\ name{\isachardoublequote}}.  These names later become the heads of parse
 | 
|
884  | 
trees; they also guide the pretty printer.  | 
|
885  | 
||
886  | 
  Productions without such parse tree names are called \emph{copy
 | 
|
887  | 
productions}. Their right-hand side must have exactly one  | 
|
888  | 
nonterminal symbol (or named token). The parser does not create a  | 
|
889  | 
new parse tree node for copy productions, but simply returns the  | 
|
890  | 
parse tree of the right-hand symbol.  | 
|
891  | 
||
892  | 
If the right-hand side of a copy production consists of a single  | 
|
893  | 
  nonterminal without any delimiters, then it is called a \emph{chain
 | 
|
894  | 
production}. Chain productions act as abbreviations: conceptually,  | 
|
895  | 
they are removed from the grammar by adding new productions.  | 
|
896  | 
Priority information attached to chain productions is ignored; only  | 
|
897  | 
  the dummy value \isa{{\isachardoublequote}{\isacharminus}{\isadigit{1}}{\isachardoublequote}} is displayed.
 | 
|
898  | 
||
| 28857 | 899  | 
  \item \isa{{\isachardoublequote}print\ modes{\isachardoublequote}} lists the alternative print modes
 | 
| 28762 | 900  | 
  provided by this grammar; see \secref{sec:print-modes}.
 | 
901  | 
||
902  | 
  \item \isa{{\isachardoublequote}parse{\isacharunderscore}rules{\isachardoublequote}} and \isa{{\isachardoublequote}print{\isacharunderscore}rules{\isachardoublequote}} relate to
 | 
|
903  | 
  syntax translations (macros); see \secref{sec:syn-trans}.
 | 
|
904  | 
||
905  | 
  \item \isa{{\isachardoublequote}parse{\isacharunderscore}ast{\isacharunderscore}translation{\isachardoublequote}} and \isa{{\isachardoublequote}print{\isacharunderscore}ast{\isacharunderscore}translation{\isachardoublequote}} list sets of constants that invoke
 | 
|
906  | 
translation functions for abstract syntax trees, which are only  | 
|
907  | 
  required in very special situations; see \secref{sec:tr-funs}.
 | 
|
908  | 
||
909  | 
  \item \isa{{\isachardoublequote}parse{\isacharunderscore}translation{\isachardoublequote}} and \isa{{\isachardoublequote}print{\isacharunderscore}translation{\isachardoublequote}}
 | 
|
910  | 
list the sets of constants that invoke regular translation  | 
|
911  | 
  functions; see \secref{sec:tr-funs}.
 | 
|
912  | 
||
913  | 
  \end{description}
 | 
|
914  | 
||
915  | 
  \end{description}%
 | 
|
916  | 
\end{isamarkuptext}%
 | 
|
917  | 
\isamarkuptrue%  | 
|
918  | 
%  | 
|
919  | 
\isadelimtheory  | 
|
920  | 
%  | 
|
921  | 
\endisadelimtheory  | 
|
922  | 
%  | 
|
923  | 
\isatagtheory  | 
|
924  | 
\isacommand{end}\isamarkupfalse%
 | 
|
925  | 
%  | 
|
926  | 
\endisatagtheory  | 
|
927  | 
{\isafoldtheory}%
 | 
|
928  | 
%  | 
|
929  | 
\isadelimtheory  | 
|
930  | 
%  | 
|
931  | 
\endisadelimtheory  | 
|
932  | 
\isanewline  | 
|
933  | 
\end{isabellebody}%
 | 
|
934  | 
%%% Local Variables:  | 
|
935  | 
%%% mode: latex  | 
|
936  | 
%%% TeX-master: "root"  | 
|
937  | 
%%% End:  |