author | ballarin |
Tue, 28 Oct 2008 11:03:07 +0100 | |
changeset 28699 | 32b6a8f12c1c |
parent 28504 | 7ad7d7d6df47 |
child 28746 | c1b151a60a66 |
permissions | -rw-r--r-- |
27043 | 1 |
(* $Id$ *) |
2 |
||
3 |
theory Document_Preparation |
|
4 |
imports Main |
|
5 |
begin |
|
6 |
||
7 |
chapter {* Document preparation \label{ch:document-prep} *} |
|
8 |
||
9 |
text {* |
|
10 |
Isabelle/Isar provides a simple document preparation system based on |
|
11 |
existing {PDF-\LaTeX} technology, with full support of hyper-links |
|
12 |
(both local references and URLs) and bookmarks. Thus the results |
|
13 |
are equally well suited for WWW browsing and as printed copies. |
|
14 |
||
15 |
\medskip Isabelle generates {\LaTeX} output as part of the run of a |
|
16 |
\emph{logic session} (see also \cite{isabelle-sys}). Getting |
|
17 |
started with a working configuration for common situations is quite |
|
18 |
easy by using the Isabelle @{verbatim mkdir} and @{verbatim make} |
|
19 |
tools. First invoke |
|
20 |
\begin{ttbox} |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
27881
diff
changeset
|
21 |
isabelle mkdir Foo |
27043 | 22 |
\end{ttbox} |
23 |
to initialize a separate directory for session @{verbatim Foo} --- |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
27881
diff
changeset
|
24 |
it is safe to experiment, since @{verbatim "isabelle mkdir"} never |
27043 | 25 |
overwrites existing files. Ensure that @{verbatim "Foo/ROOT.ML"} |
26 |
holds ML commands to load all theories required for this session; |
|
27 |
furthermore @{verbatim "Foo/document/root.tex"} should include any |
|
28 |
special {\LaTeX} macro packages required for your document (the |
|
29 |
default is usually sufficient as a start). |
|
30 |
||
31 |
The session is controlled by a separate @{verbatim IsaMakefile} |
|
32 |
(with crude source dependencies by default). This file is located |
|
33 |
one level up from the @{verbatim Foo} directory location. Now |
|
34 |
invoke |
|
35 |
\begin{ttbox} |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
27881
diff
changeset
|
36 |
isabelle make Foo |
27043 | 37 |
\end{ttbox} |
38 |
to run the @{verbatim Foo} session, with browser information and |
|
39 |
document preparation enabled. Unless any errors are reported by |
|
40 |
Isabelle or {\LaTeX}, the output will appear inside the directory |
|
41 |
@{verbatim ISABELLE_BROWSER_INFO}, as reported by the batch job in |
|
42 |
verbose mode. |
|
43 |
||
44 |
\medskip You may also consider to tune the @{verbatim usedir} |
|
45 |
options in @{verbatim IsaMakefile}, for example to change the output |
|
46 |
format from @{verbatim pdf} to @{verbatim dvi}, or activate the |
|
47 |
@{verbatim "-D"} option to retain a second copy of the generated |
|
48 |
{\LaTeX} sources. |
|
49 |
||
50 |
\medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys} |
|
51 |
for further details on Isabelle logic sessions and theory |
|
52 |
presentation. The Isabelle/HOL tutorial \cite{isabelle-hol-book} |
|
53 |
also covers theory presentation issues. |
|
54 |
*} |
|
55 |
||
56 |
||
57 |
section {* Markup commands \label{sec:markup} *} |
|
58 |
||
59 |
text {* |
|
60 |
\begin{matharray}{rcl} |
|
27049 | 61 |
@{command_def "header"} & : & \isarkeep{toplevel} \\[0.5ex] |
27043 | 62 |
@{command_def "chapter"} & : & \isarkeep{local{\dsh}theory} \\ |
63 |
@{command_def "section"} & : & \isarkeep{local{\dsh}theory} \\ |
|
64 |
@{command_def "subsection"} & : & \isarkeep{local{\dsh}theory} \\ |
|
65 |
@{command_def "subsubsection"} & : & \isarkeep{local{\dsh}theory} \\ |
|
66 |
@{command_def "text"} & : & \isarkeep{local{\dsh}theory} \\ |
|
67 |
@{command_def "text_raw"} & : & \isarkeep{local{\dsh}theory} \\[0.5ex] |
|
68 |
@{command_def "sect"} & : & \isartrans{proof}{proof} \\ |
|
69 |
@{command_def "subsect"} & : & \isartrans{proof}{proof} \\ |
|
70 |
@{command_def "subsubsect"} & : & \isartrans{proof}{proof} \\ |
|
71 |
@{command_def "txt"} & : & \isartrans{proof}{proof} \\ |
|
72 |
@{command_def "txt_raw"} & : & \isartrans{proof}{proof} \\ |
|
73 |
\end{matharray} |
|
74 |
||
75 |
Apart from formal comments (see \secref{sec:comments}), markup |
|
76 |
commands provide a structured way to insert text into the document |
|
77 |
generated from a theory (see \cite{isabelle-sys} for more |
|
78 |
information on Isabelle's document preparation tools). |
|
79 |
||
80 |
\begin{rail} |
|
81 |
('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text') target? text |
|
82 |
; |
|
27049 | 83 |
('header' | 'text\_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt\_raw') text |
27043 | 84 |
; |
85 |
\end{rail} |
|
86 |
||
87 |
\begin{descr} |
|
88 |
||
27049 | 89 |
\item [@{command "header"}~@{text "text"}] provides plain text |
90 |
markup just preceding the formal beginning of a theory. In actual |
|
91 |
document preparation the corresponding {\LaTeX} macro @{verbatim |
|
92 |
"\\isamarkupheader"} may be redefined to produce chapter or section |
|
27053 | 93 |
headings. |
27049 | 94 |
|
27043 | 95 |
\item [@{command "chapter"}, @{command "section"}, @{command |
96 |
"subsection"}, and @{command "subsubsection"}] mark chapter and |
|
27053 | 97 |
section headings. The corresponding {\LaTeX} macros are @{verbatim |
98 |
"\\isamarkupchapter"}, @{verbatim "\\isamarkupsection"} etc. |
|
27043 | 99 |
|
100 |
\item [@{command "text"} and @{command "txt"}] specify paragraphs of |
|
101 |
plain text. |
|
102 |
||
103 |
\item [@{command "text_raw"} and @{command "txt_raw"}] insert |
|
104 |
{\LaTeX} source into the output, without additional markup. Thus |
|
105 |
the full range of document manipulations becomes available. |
|
106 |
||
107 |
\end{descr} |
|
108 |
||
109 |
The @{text "text"} argument of these markup commands (except for |
|
110 |
@{command "text_raw"}) may contain references to formal entities |
|
111 |
(``antiquotations'', see also \secref{sec:antiq}). These are |
|
112 |
interpreted in the present theory context, or the named @{text |
|
113 |
"target"}. |
|
114 |
||
115 |
Any of these markup elements corresponds to a {\LaTeX} command with |
|
116 |
the name prefixed by @{verbatim "\\isamarkup"}. For the sectioning |
|
117 |
commands this is a plain macro with a single argument, e.g.\ |
|
118 |
@{verbatim "\\isamarkupchapter{"}@{text "\<dots>"}@{verbatim "}"} for |
|
119 |
@{command "chapter"}. The @{command "text"} markup results in a |
|
120 |
{\LaTeX} environment @{verbatim "\\begin{isamarkuptext}"} @{text |
|
121 |
"\<dots>"} @{verbatim "\\end{isamarkuptext}"}, while @{command "text_raw"} |
|
122 |
causes the text to be inserted directly into the {\LaTeX} source. |
|
123 |
||
124 |
\medskip The proof markup commands closely resemble those for theory |
|
125 |
specifications, but have a different formal status and produce |
|
126 |
different {\LaTeX} macros. Also note that the @{command_ref |
|
127 |
"header"} declaration (see \secref{sec:begin-thy}) admits to insert |
|
128 |
section markup just preceding the actual theory definition. |
|
129 |
*} |
|
130 |
||
131 |
||
132 |
section {* Antiquotations \label{sec:antiq} *} |
|
133 |
||
134 |
text {* |
|
135 |
\begin{matharray}{rcl} |
|
136 |
@{antiquotation_def "theory"} & : & \isarantiq \\ |
|
137 |
@{antiquotation_def "thm"} & : & \isarantiq \\ |
|
27453 | 138 |
@{antiquotation_def "lemma"} & : & \isarantiq \\ |
27043 | 139 |
@{antiquotation_def "prop"} & : & \isarantiq \\ |
140 |
@{antiquotation_def "term"} & : & \isarantiq \\ |
|
141 |
@{antiquotation_def const} & : & \isarantiq \\ |
|
142 |
@{antiquotation_def abbrev} & : & \isarantiq \\ |
|
143 |
@{antiquotation_def typeof} & : & \isarantiq \\ |
|
144 |
@{antiquotation_def typ} & : & \isarantiq \\ |
|
145 |
@{antiquotation_def thm_style} & : & \isarantiq \\ |
|
146 |
@{antiquotation_def term_style} & : & \isarantiq \\ |
|
147 |
@{antiquotation_def "text"} & : & \isarantiq \\ |
|
148 |
@{antiquotation_def goals} & : & \isarantiq \\ |
|
149 |
@{antiquotation_def subgoals} & : & \isarantiq \\ |
|
150 |
@{antiquotation_def prf} & : & \isarantiq \\ |
|
151 |
@{antiquotation_def full_prf} & : & \isarantiq \\ |
|
152 |
@{antiquotation_def ML} & : & \isarantiq \\ |
|
153 |
@{antiquotation_def ML_type} & : & \isarantiq \\ |
|
154 |
@{antiquotation_def ML_struct} & : & \isarantiq \\ |
|
155 |
\end{matharray} |
|
156 |
||
157 |
The text body of formal comments (see also \secref{sec:comments}) |
|
158 |
may contain antiquotations of logical entities, such as theorems, |
|
159 |
terms and types, which are to be presented in the final output |
|
27053 | 160 |
produced by the Isabelle document preparation system. |
27043 | 161 |
|
27881
f0d543629376
[source=false] for quoted antiquotation avoids quote-escapes in output;
wenzelm
parents:
27453
diff
changeset
|
162 |
Thus embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}'' |
27043 | 163 |
within a text block would cause |
164 |
\isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} to appear in the final {\LaTeX} document. Also note that theorem |
|
165 |
antiquotations may involve attributes as well. For example, |
|
166 |
@{text "@{thm sym [no_vars]}"} would print the theorem's |
|
167 |
statement where all schematic variables have been replaced by fixed |
|
168 |
ones, which are easier to read. |
|
169 |
||
170 |
\begin{rail} |
|
171 |
atsign lbrace antiquotation rbrace |
|
172 |
; |
|
173 |
||
174 |
antiquotation: |
|
175 |
'theory' options name | |
|
176 |
'thm' options thmrefs | |
|
27453 | 177 |
'lemma' options prop 'by' method | |
27043 | 178 |
'prop' options prop | |
179 |
'term' options term | |
|
180 |
'const' options term | |
|
181 |
'abbrev' options term | |
|
182 |
'typeof' options term | |
|
183 |
'typ' options type | |
|
184 |
'thm\_style' options name thmref | |
|
185 |
'term\_style' options name term | |
|
186 |
'text' options name | |
|
187 |
'goals' options | |
|
188 |
'subgoals' options | |
|
189 |
'prf' options thmrefs | |
|
190 |
'full\_prf' options thmrefs | |
|
191 |
'ML' options name | |
|
192 |
'ML\_type' options name | |
|
193 |
'ML\_struct' options name |
|
194 |
; |
|
195 |
options: '[' (option * ',') ']' |
|
196 |
; |
|
197 |
option: name | name '=' name |
|
198 |
; |
|
199 |
\end{rail} |
|
200 |
||
201 |
Note that the syntax of antiquotations may \emph{not} include source |
|
202 |
comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} or verbatim |
|
203 |
text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim |
|
204 |
"*"}@{verbatim "}"}. |
|
205 |
||
206 |
\begin{descr} |
|
207 |
||
208 |
\item [@{text "@{theory A}"}] prints the name @{text "A"}, which is |
|
209 |
guaranteed to refer to a valid ancestor theory in the current |
|
210 |
context. |
|
211 |
||
212 |
\item [@{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"}] prints theorems |
|
213 |
@{text "a\<^sub>1 \<dots> a\<^sub>n"}. Note that attribute specifications |
|
214 |
may be included as well (see also \secref{sec:syn-att}); the |
|
215 |
@{attribute_ref no_vars} rule (see \secref{sec:misc-meth-att}) would |
|
216 |
be particularly useful to suppress printing of schematic variables. |
|
217 |
||
218 |
\item [@{text "@{prop \<phi>}"}] prints a well-typed proposition @{text |
|
219 |
"\<phi>"}. |
|
220 |
||
27453 | 221 |
\item [@{text "@{lemma \<phi> by m}"}] asserts a well-typed proposition @{text |
222 |
"\<phi>"} to be provable by method @{text m} and prints @{text "\<phi>"}. |
|
223 |
||
27043 | 224 |
\item [@{text "@{term t}"}] prints a well-typed term @{text "t"}. |
225 |
||
226 |
\item [@{text "@{const c}"}] prints a logical or syntactic constant |
|
227 |
@{text "c"}. |
|
228 |
||
229 |
\item [@{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"}] prints a constant |
|
230 |
abbreviation @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in |
|
231 |
the current context. |
|
232 |
||
233 |
\item [@{text "@{typeof t}"}] prints the type of a well-typed term |
|
234 |
@{text "t"}. |
|
235 |
||
236 |
\item [@{text "@{typ \<tau>}"}] prints a well-formed type @{text "\<tau>"}. |
|
237 |
||
238 |
\item [@{text "@{thm_style s a}"}] prints theorem @{text a}, |
|
239 |
previously applying a style @{text s} to it (see below). |
|
240 |
||
241 |
\item [@{text "@{term_style s t}"}] prints a well-typed term @{text |
|
242 |
t} after applying a style @{text s} to it (see below). |
|
243 |
||
244 |
\item [@{text "@{text s}"}] prints uninterpreted source text @{text |
|
245 |
s}. This is particularly useful to print portions of text according |
|
246 |
to the Isabelle {\LaTeX} output style, without demanding |
|
247 |
well-formedness (e.g.\ small pieces of terms that should not be |
|
248 |
parsed or type-checked yet). |
|
249 |
||
250 |
\item [@{text "@{goals}"}] prints the current \emph{dynamic} goal |
|
251 |
state. This is mainly for support of tactic-emulation scripts |
|
252 |
within Isar --- presentation of goal states does not conform to |
|
253 |
actual human-readable proof documents. |
|
254 |
||
255 |
Please do not include goal states into document output unless you |
|
256 |
really know what you are doing! |
|
257 |
||
258 |
\item [@{text "@{subgoals}"}] is similar to @{text "@{goals}"}, but |
|
259 |
does not print the main goal. |
|
260 |
||
261 |
\item [@{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"}] prints the (compact) |
|
262 |
proof terms corresponding to the theorems @{text "a\<^sub>1 \<dots> |
|
263 |
a\<^sub>n"}. Note that this requires proof terms to be switched on |
|
264 |
for the current object logic (see the ``Proof terms'' section of the |
|
265 |
Isabelle reference manual for information on how to do this). |
|
266 |
||
267 |
\item [@{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"}] is like @{text |
|
268 |
"@{prf a\<^sub>1 \<dots> a\<^sub>n}"}, but displays the full proof terms, |
|
269 |
i.e.\ also displays information omitted in the compact proof term, |
|
270 |
which is denoted by ``@{text _}'' placeholders there. |
|
271 |
||
272 |
\item [@{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text |
|
273 |
"@{ML_struct s}"}] check text @{text s} as ML value, type, and |
|
274 |
structure, respectively. The source is displayed verbatim. |
|
275 |
||
276 |
\end{descr} |
|
277 |
||
278 |
\medskip The following standard styles for use with @{text |
|
279 |
thm_style} and @{text term_style} are available: |
|
280 |
||
281 |
\begin{descr} |
|
282 |
||
283 |
\item [@{text lhs}] extracts the first argument of any application |
|
284 |
form with at least two arguments -- typically meta-level or |
|
285 |
object-level equality, or any other binary relation. |
|
286 |
||
287 |
\item [@{text rhs}] is like @{text lhs}, but extracts the second |
|
288 |
argument. |
|
289 |
||
290 |
\item [@{text "concl"}] extracts the conclusion @{text C} from a rule |
|
291 |
in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}. |
|
292 |
||
293 |
\item [@{text "prem1"}, \dots, @{text "prem9"}] extract premise |
|
294 |
number @{text "1, \<dots>, 9"}, respectively, from from a rule in |
|
295 |
Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"} |
|
296 |
||
297 |
\end{descr} |
|
298 |
||
299 |
\medskip |
|
300 |
The following options are available to tune the output. Note that most of |
|
301 |
these coincide with ML flags of the same names (see also \cite{isabelle-ref}). |
|
302 |
||
303 |
\begin{descr} |
|
304 |
||
305 |
\item[@{text "show_types = bool"} and @{text "show_sorts = bool"}] |
|
306 |
control printing of explicit type and sort constraints. |
|
307 |
||
308 |
\item[@{text "show_structs = bool"}] controls printing of implicit |
|
309 |
structures. |
|
310 |
||
311 |
\item[@{text "long_names = bool"}] forces names of types and |
|
312 |
constants etc.\ to be printed in their fully qualified internal |
|
313 |
form. |
|
314 |
||
315 |
\item[@{text "short_names = bool"}] forces names of types and |
|
316 |
constants etc.\ to be printed unqualified. Note that internalizing |
|
317 |
the output again in the current context may well yield a different |
|
318 |
result. |
|
319 |
||
320 |
\item[@{text "unique_names = bool"}] determines whether the printed |
|
321 |
version of qualified names should be made sufficiently long to avoid |
|
322 |
overlap with names declared further back. Set to @{text false} for |
|
323 |
more concise output. |
|
324 |
||
325 |
\item[@{text "eta_contract = bool"}] prints terms in @{text |
|
326 |
\<eta>}-contracted form. |
|
327 |
||
328 |
\item[@{text "display = bool"}] indicates if the text is to be |
|
329 |
output as multi-line ``display material'', rather than a small piece |
|
330 |
of text without line breaks (which is the default). |
|
331 |
||
332 |
\item[@{text "break = bool"}] controls line breaks in non-display |
|
333 |
material. |
|
334 |
||
335 |
\item[@{text "quotes = bool"}] indicates if the output should be |
|
336 |
enclosed in double quotes. |
|
337 |
||
338 |
\item[@{text "mode = name"}] adds @{text name} to the print mode to |
|
339 |
be used for presentation (see also \cite{isabelle-ref}). Note that |
|
340 |
the standard setup for {\LaTeX} output is already present by |
|
341 |
default, including the modes @{text latex} and @{text xsymbols}. |
|
342 |
||
343 |
\item[@{text "margin = nat"} and @{text "indent = nat"}] change the |
|
344 |
margin or indentation for pretty printing of display material. |
|
345 |
||
346 |
\item[@{text "source = bool"}] prints the source text of the |
|
347 |
antiquotation arguments, rather than the actual value. Note that |
|
348 |
this does not affect well-formedness checks of @{antiquotation |
|
349 |
"thm"}, @{antiquotation "term"}, etc. (only the @{antiquotation |
|
350 |
"text"} antiquotation admits arbitrary output). |
|
351 |
||
352 |
\item[@{text "goals_limit = nat"}] determines the maximum number of |
|
353 |
goals to be printed. |
|
354 |
||
355 |
\item[@{text "locale = name"}] specifies an alternative locale |
|
356 |
context used for evaluating and printing the subsequent argument. |
|
357 |
||
358 |
\end{descr} |
|
359 |
||
360 |
For boolean flags, ``@{text "name = true"}'' may be abbreviated as |
|
361 |
``@{text name}''. All of the above flags are disabled by default, |
|
362 |
unless changed from ML. |
|
363 |
||
364 |
\medskip Note that antiquotations do not only spare the author from |
|
365 |
tedious typing of logical entities, but also achieve some degree of |
|
366 |
consistency-checking of informal explanations with formal |
|
367 |
developments: well-formedness of terms and types with respect to the |
|
368 |
current theory or proof context is ensured here. |
|
369 |
*} |
|
370 |
||
371 |
||
372 |
section {* Tagged commands \label{sec:tags} *} |
|
373 |
||
374 |
text {* |
|
375 |
Each Isabelle/Isar command may be decorated by presentation tags: |
|
376 |
||
377 |
\indexouternonterm{tags} |
|
378 |
\begin{rail} |
|
379 |
tags: ( tag * ) |
|
380 |
; |
|
381 |
tag: '\%' (ident | string) |
|
382 |
\end{rail} |
|
383 |
||
384 |
The tags @{text "theory"}, @{text "proof"}, @{text "ML"} are already |
|
385 |
pre-declared for certain classes of commands: |
|
386 |
||
387 |
\medskip |
|
388 |
||
389 |
\begin{tabular}{ll} |
|
390 |
@{text "theory"} & theory begin/end \\ |
|
391 |
@{text "proof"} & all proof commands \\ |
|
392 |
@{text "ML"} & all commands involving ML code \\ |
|
393 |
\end{tabular} |
|
394 |
||
395 |
\medskip The Isabelle document preparation system (see also |
|
396 |
\cite{isabelle-sys}) allows tagged command regions to be presented |
|
397 |
specifically, e.g.\ to fold proof texts, or drop parts of the text |
|
398 |
completely. |
|
399 |
||
400 |
For example ``@{command "by"}~@{text "%invisible auto"}'' would |
|
401 |
cause that piece of proof to be treated as @{text invisible} instead |
|
402 |
of @{text "proof"} (the default), which may be either show or hidden |
|
403 |
depending on the document setup. In contrast, ``@{command |
|
404 |
"by"}~@{text "%visible auto"}'' would force this text to be shown |
|
405 |
invariably. |
|
406 |
||
407 |
Explicit tag specifications within a proof apply to all subsequent |
|
408 |
commands of the same level of nesting. For example, ``@{command |
|
409 |
"proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' would force the |
|
410 |
whole sub-proof to be typeset as @{text visible} (unless some of its |
|
411 |
parts are tagged differently). |
|
412 |
*} |
|
413 |
||
414 |
||
415 |
section {* Draft presentation *} |
|
416 |
||
417 |
text {* |
|
418 |
\begin{matharray}{rcl} |
|
419 |
@{command_def "display_drafts"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\ |
|
420 |
@{command_def "print_drafts"}@{text "\<^sup>*"} & : & \isarkeep{\cdot} \\ |
|
421 |
\end{matharray} |
|
422 |
||
423 |
\begin{rail} |
|
424 |
('display\_drafts' | 'print\_drafts') (name +) |
|
425 |
; |
|
426 |
\end{rail} |
|
427 |
||
428 |
\begin{descr} |
|
429 |
||
430 |
\item [@{command "display_drafts"}~@{text paths} and @{command |
|
431 |
"print_drafts"}~@{text paths}] perform simple output of a given list |
|
432 |
of raw source files. Only those symbols that do not require |
|
433 |
additional {\LaTeX} packages are displayed properly, everything else |
|
434 |
is left verbatim. |
|
435 |
||
436 |
\end{descr} |
|
437 |
*} |
|
438 |
||
439 |
end |