| author | nipkow |
| Fri, 30 Oct 2015 20:01:05 +0100 | |
| changeset 61525 | 87244a9cfe40 |
| parent 61504 | a7ae3ef886a9 |
| child 61595 | 3591274c607e |
| permissions | -rw-r--r-- |
| 27043 | 1 |
theory Document_Preparation |
| 42651 | 2 |
imports Base Main |
| 27043 | 3 |
begin |
4 |
||
| 58618 | 5 |
chapter \<open>Document preparation \label{ch:document-prep}\<close>
|
| 27043 | 6 |
|
| 58618 | 7 |
text \<open>Isabelle/Isar provides a simple document preparation system |
| 54346 | 8 |
based on {PDF-\LaTeX}, with support for hyperlinks and bookmarks
|
9 |
within that format. This allows to produce papers, books, theses |
|
10 |
etc.\ from Isabelle theory sources. |
|
| 27043 | 11 |
|
| 61477 | 12 |
{\LaTeX} output is generated while processing a \<^emph>\<open>session\<close> in
|
13 |
batch mode, as explained in the \<^emph>\<open>The Isabelle System Manual\<close> |
|
| 60270 | 14 |
@{cite "isabelle-system"}. The main Isabelle tools to get started with
|
| 54346 | 15 |
document preparation are @{tool_ref mkroot} and @{tool_ref build}.
|
| 27043 | 16 |
|
| 58552 | 17 |
The classic Isabelle/HOL tutorial @{cite "isabelle-hol-book"} also
|
| 58618 | 18 |
explains some aspects of theory presentation.\<close> |
| 27043 | 19 |
|
20 |
||
| 58618 | 21 |
section \<open>Markup commands \label{sec:markup}\<close>
|
| 27043 | 22 |
|
| 58618 | 23 |
text \<open> |
| 27043 | 24 |
\begin{matharray}{rcl}
|
| 61493 | 25 |
@{command_def "chapter"} & : & \<open>any \<rightarrow> any\<close> \\
|
26 |
@{command_def "section"} & : & \<open>any \<rightarrow> any\<close> \\
|
|
27 |
@{command_def "subsection"} & : & \<open>any \<rightarrow> any\<close> \\
|
|
28 |
@{command_def "subsubsection"} & : & \<open>any \<rightarrow> any\<close> \\
|
|
29 |
@{command_def "paragraph"} & : & \<open>any \<rightarrow> any\<close> \\
|
|
30 |
@{command_def "subparagraph"} & : & \<open>any \<rightarrow> any\<close> \\
|
|
31 |
@{command_def "text"} & : & \<open>any \<rightarrow> any\<close> \\
|
|
32 |
@{command_def "txt"} & : & \<open>any \<rightarrow> any\<close> \\
|
|
33 |
@{command_def "text_raw"} & : & \<open>any \<rightarrow> any\<close> \\
|
|
| 27043 | 34 |
\end{matharray}
|
35 |
||
| 28747 | 36 |
Markup commands provide a structured way to insert text into the |
37 |
document generated from a theory. Each markup command takes a |
|
38 |
single @{syntax text} argument, which is passed as argument to a
|
|
39 |
corresponding {\LaTeX} macro. The default macros provided by
|
|
|
40800
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents:
40255
diff
changeset
|
40 |
@{file "~~/lib/texinputs/isabelle.sty"} can be redefined according
|
| 28747 | 41 |
to the needs of the underlying document and {\LaTeX} styles.
|
42 |
||
43 |
Note that formal comments (\secref{sec:comments}) are similar to
|
|
44 |
markup commands, but have a different status within Isabelle/Isar |
|
45 |
syntax. |
|
| 27043 | 46 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
47 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
48 |
(@@{command chapter} | @@{command section} | @@{command subsection} |
|
| 61463 | 49 |
@@{command subsubsection} | @@{command paragraph} | @@{command subparagraph} |
|
50 |
@@{command text} | @@{command txt} | @@{command text_raw}) @{syntax text}
|
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
51 |
\<close>} |
| 27043 | 52 |
|
| 61463 | 53 |
\<^descr> @{command chapter}, @{command section}, @{command subsection} etc.\ mark
|
54 |
section headings within the theory source. This works in any context, even |
|
55 |
before the initial @{command theory} command. The corresponding {\LaTeX}
|
|
| 61503 | 56 |
macros are \<^verbatim>\<open>\isamarkupchapter\<close>, \<^verbatim>\<open>\isamarkupsection\<close>, \<^verbatim>\<open>\isamarkupsubsection\<close> etc.\ |
| 27043 | 57 |
|
| 61439 | 58 |
\<^descr> @{command text} and @{command txt} specify paragraphs of plain text.
|
| 61503 | 59 |
This corresponds to a {\LaTeX} environment \<^verbatim>\<open>\begin{isamarkuptext}\<close> \<open>\<dots>\<close>
|
60 |
\<^verbatim>\<open>\end{isamarkuptext}\<close> etc.
|
|
| 27043 | 61 |
|
| 61463 | 62 |
\<^descr> @{command text_raw} is similar to @{command text}, but without
|
63 |
any surrounding markup environment. This allows to inject arbitrary |
|
64 |
{\LaTeX} source into the generated document.
|
|
| 27043 | 65 |
|
66 |
||
| 61463 | 67 |
All text passed to any of the above markup commands may refer to formal |
| 61477 | 68 |
entities via \<^emph>\<open>document antiquotations\<close>, see also \secref{sec:antiq}.
|
| 61463 | 69 |
These are interpreted in the present theory or proof context. |
| 27043 | 70 |
|
| 61421 | 71 |
\<^medskip> |
72 |
The proof markup commands closely resemble those for theory |
|
| 27043 | 73 |
specifications, but have a different formal status and produce |
|
58868
c5e1cce7ace3
uniform heading commands work in any context, even in theory header;
wenzelm
parents:
58725
diff
changeset
|
74 |
different {\LaTeX} macros.
|
| 58618 | 75 |
\<close> |
| 27043 | 76 |
|
77 |
||
| 60286 | 78 |
section \<open>Document antiquotations \label{sec:antiq}\<close>
|
| 27043 | 79 |
|
| 58618 | 80 |
text \<open> |
| 27043 | 81 |
\begin{matharray}{rcl}
|
| 61493 | 82 |
@{antiquotation_def "theory"} & : & \<open>antiquotation\<close> \\
|
83 |
@{antiquotation_def "thm"} & : & \<open>antiquotation\<close> \\
|
|
84 |
@{antiquotation_def "lemma"} & : & \<open>antiquotation\<close> \\
|
|
85 |
@{antiquotation_def "prop"} & : & \<open>antiquotation\<close> \\
|
|
86 |
@{antiquotation_def "term"} & : & \<open>antiquotation\<close> \\
|
|
87 |
@{antiquotation_def term_type} & : & \<open>antiquotation\<close> \\
|
|
88 |
@{antiquotation_def typeof} & : & \<open>antiquotation\<close> \\
|
|
89 |
@{antiquotation_def const} & : & \<open>antiquotation\<close> \\
|
|
90 |
@{antiquotation_def abbrev} & : & \<open>antiquotation\<close> \\
|
|
91 |
@{antiquotation_def typ} & : & \<open>antiquotation\<close> \\
|
|
92 |
@{antiquotation_def type} & : & \<open>antiquotation\<close> \\
|
|
93 |
@{antiquotation_def class} & : & \<open>antiquotation\<close> \\
|
|
94 |
@{antiquotation_def "text"} & : & \<open>antiquotation\<close> \\
|
|
95 |
@{antiquotation_def goals} & : & \<open>antiquotation\<close> \\
|
|
96 |
@{antiquotation_def subgoals} & : & \<open>antiquotation\<close> \\
|
|
97 |
@{antiquotation_def prf} & : & \<open>antiquotation\<close> \\
|
|
98 |
@{antiquotation_def full_prf} & : & \<open>antiquotation\<close> \\
|
|
99 |
@{antiquotation_def ML} & : & \<open>antiquotation\<close> \\
|
|
100 |
@{antiquotation_def ML_op} & : & \<open>antiquotation\<close> \\
|
|
101 |
@{antiquotation_def ML_type} & : & \<open>antiquotation\<close> \\
|
|
102 |
@{antiquotation_def ML_structure} & : & \<open>antiquotation\<close> \\
|
|
103 |
@{antiquotation_def ML_functor} & : & \<open>antiquotation\<close> \\
|
|
104 |
@{antiquotation_def emph} & : & \<open>antiquotation\<close> \\
|
|
105 |
@{antiquotation_def bold} & : & \<open>antiquotation\<close> \\
|
|
106 |
@{antiquotation_def verbatim} & : & \<open>antiquotation\<close> \\
|
|
107 |
@{antiquotation_def "file"} & : & \<open>antiquotation\<close> \\
|
|
108 |
@{antiquotation_def "url"} & : & \<open>antiquotation\<close> \\
|
|
109 |
@{antiquotation_def "cite"} & : & \<open>antiquotation\<close> \\
|
|
| 61494 | 110 |
@{command_def "print_antiquotations"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
|
| 27043 | 111 |
\end{matharray}
|
112 |
||
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
113 |
The overall content of an Isabelle/Isar theory may alternate between |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
114 |
formal and informal text. The main body consists of formal |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
115 |
specification and proof commands, interspersed with markup commands |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
116 |
(\secref{sec:markup}) or document comments (\secref{sec:comments}).
|
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
117 |
The argument of markup commands quotes informal text to be printed |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
118 |
in the resulting document, but may again refer to formal entities |
| 61477 | 119 |
via \<^emph>\<open>document antiquotations\<close>. |
| 27043 | 120 |
|
| 61503 | 121 |
For example, embedding \<^verbatim>\<open>@{term [show_types] "f x = a + x"}\<close>
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
122 |
within a text block makes |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
123 |
\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} appear in the final {\LaTeX} document.
|
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
124 |
|
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
125 |
Antiquotations usually spare the author tedious typing of logical |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
126 |
entities in full detail. Even more importantly, some degree of |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
127 |
consistency-checking between the main body of formal text and its |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
128 |
informal explanation is achieved, since terms and types appearing in |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
129 |
antiquotations are checked within the current theory or proof |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
130 |
context. |
| 27043 | 131 |
|
| 61503 | 132 |
\<^medskip> Antiquotations are in general written as \<^verbatim>\<open>@{\<close>\<open>name\<close>~\<^verbatim>\<open>[\<close>\<open>options\<close>\<^verbatim>\<open>]\<close>~\<open>arguments\<close>\<^verbatim>\<open>}\<close>.
|
133 |
The short form \<^verbatim>\<open>\\<close>\<^verbatim>\<open><^\<close>\<open>name\<close>\<^verbatim>\<open>>\<close>\<open>\<open>argument_content\<close>\<close> (without |
|
134 |
surrounding \<^verbatim>\<open>@{\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close>) works for a single
|
|
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
135 |
argument that is a cartouche. |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
136 |
|
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
137 |
Omitting the control symbol is also possible: a cartouche without special |
| 61504 | 138 |
decoration is equivalent to \<^verbatim>\<open>\<^cartouche>\<close>\<open>\<open>argument_content\<close>\<close>, which |
| 61503 | 139 |
is equivalent to \<^verbatim>\<open>@{cartouche\<close>~\<open>\<open>argument_content\<close>\<close>\<^verbatim>\<open>}\<close>. The
|
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
140 |
special name @{antiquotation_def cartouche} is defined in the context:
|
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
141 |
Isabelle/Pure introduces that as an alias to @{antiquotation_ref text}
|
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
142 |
(see below). Consequently, \<open>\<open>foo_bar + baz \<le> bazar\<close>\<close> prints literal |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
143 |
quasi-formal text (unchecked). |
| 61472 | 144 |
|
| 61474 | 145 |
\begingroup |
146 |
\def\isasymcontrolstart{\isatt{\isacharbackslash\isacharless\isacharcircum}}
|
|
|
59917
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59783
diff
changeset
|
147 |
@{rail \<open>
|
| 61472 | 148 |
@{syntax_def antiquotation}:
|
| 61474 | 149 |
'@{' antiquotation_body '}' |
|
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
150 |
'\<controlstart>' @{syntax_ref name} '>' @{syntax_ref cartouche} |
|
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
151 |
@{syntax_ref cartouche}
|
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
152 |
; |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
153 |
options: '[' (option * ',') ']' |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
154 |
; |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
155 |
option: @{syntax name} | @{syntax name} '=' @{syntax name}
|
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
156 |
; |
|
59917
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59783
diff
changeset
|
157 |
\<close>} |
| 61474 | 158 |
\endgroup |
|
59917
9830c944670f
more uniform "verbose" option to print name space;
wenzelm
parents:
59783
diff
changeset
|
159 |
|
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
160 |
Note that the syntax of antiquotations may \<^emph>\<open>not\<close> include source |
| 61503 | 161 |
comments \<^verbatim>\<open>(*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*)\<close> nor verbatim text \<^verbatim>\<open>{*\<close>~\<open>\<dots>\<close>~\<^verbatim>\<open>*}\<close>.
|
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
162 |
|
| 43618 | 163 |
%% FIXME less monolithic presentation, move to individual sections!? |
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
164 |
@{rail \<open>
|
| 61472 | 165 |
@{syntax_def antiquotation_body}:
|
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
166 |
(@@{antiquotation text} | @@{antiquotation cartouche}) options @{syntax text} |
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
167 |
@@{antiquotation theory} options @{syntax name} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
168 |
@@{antiquotation thm} options styles @{syntax thmrefs} |
|
| 42617 | 169 |
@@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? |
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
170 |
@@{antiquotation prop} options styles @{syntax prop} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
171 |
@@{antiquotation term} options styles @{syntax term} |
|
| 43618 | 172 |
@@{antiquotation (HOL) value} options styles @{syntax term} |
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
173 |
@@{antiquotation term_type} options styles @{syntax term} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
174 |
@@{antiquotation typeof} options styles @{syntax term} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
175 |
@@{antiquotation const} options @{syntax term} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
176 |
@@{antiquotation abbrev} options @{syntax term} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
177 |
@@{antiquotation typ} options @{syntax type} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
178 |
@@{antiquotation type} options @{syntax name} |
|
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
179 |
@@{antiquotation class} options @{syntax name}
|
| 46261 | 180 |
; |
181 |
@{syntax antiquotation}:
|
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
182 |
@@{antiquotation goals} options |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
183 |
@@{antiquotation subgoals} options |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
184 |
@@{antiquotation prf} options @{syntax thmrefs} |
|
|
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
185 |
@@{antiquotation full_prf} options @{syntax thmrefs} |
|
|
58069
0255436b3d85
more liberal embedded "text", which includes cartouches;
wenzelm
parents:
56594
diff
changeset
|
186 |
@@{antiquotation ML} options @{syntax text} |
|
|
0255436b3d85
more liberal embedded "text", which includes cartouches;
wenzelm
parents:
56594
diff
changeset
|
187 |
@@{antiquotation ML_op} options @{syntax text} |
|
|
0255436b3d85
more liberal embedded "text", which includes cartouches;
wenzelm
parents:
56594
diff
changeset
|
188 |
@@{antiquotation ML_type} options @{syntax text} |
|
|
0255436b3d85
more liberal embedded "text", which includes cartouches;
wenzelm
parents:
56594
diff
changeset
|
189 |
@@{antiquotation ML_structure} options @{syntax text} |
|
|
0255436b3d85
more liberal embedded "text", which includes cartouches;
wenzelm
parents:
56594
diff
changeset
|
190 |
@@{antiquotation ML_functor} options @{syntax text} |
|
|
61473
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
191 |
@@{antiquotation emph} options @{syntax text} |
|
|
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
192 |
@@{antiquotation bold} options @{syntax text} |
|
|
58716
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents:
58618
diff
changeset
|
193 |
@@{antiquotation verbatim} options @{syntax text} |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
194 |
@@{antiquotation "file"} options @{syntax name} |
|
| 54705 | 195 |
@@{antiquotation file_unchecked} options @{syntax name} |
|
| 58593 | 196 |
@@{antiquotation url} options @{syntax name} |
|
197 |
@@{antiquotation cite} options @{syntax cartouche}? (@{syntax name} + @'and')
|
|
| 27043 | 198 |
; |
| 32892 | 199 |
styles: '(' (style + ',') ')'
|
| 32891 | 200 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
201 |
style: (@{syntax name} +)
|
|
61473
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
202 |
; |
|
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
203 |
@@{command print_antiquotations} ('!'?)
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
204 |
\<close>} |
| 27043 | 205 |
|
| 61493 | 206 |
\<^descr> \<open>@{text s}\<close> prints uninterpreted source text \<open>s\<close>. This is particularly useful to print portions of text according
|
|
61491
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
207 |
to the Isabelle document style, without demanding well-formedness, |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
208 |
e.g.\ small pieces of terms that should not be parsed or |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
209 |
type-checked yet. |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
210 |
|
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
211 |
It is also possible to write this in the short form \<open>\<open>s\<close>\<close> without any |
|
97261e6c1d42
another antiquotation short form: undecorated cartouche as alias for @{text};
wenzelm
parents:
61479
diff
changeset
|
212 |
further decoration. |
| 27043 | 213 |
|
| 61493 | 214 |
\<^descr> \<open>@{theory A}\<close> prints the name \<open>A\<close>, which is
|
| 27043 | 215 |
guaranteed to refer to a valid ancestor theory in the current |
216 |
context. |
|
217 |
||
| 61493 | 218 |
\<^descr> \<open>@{thm a\<^sub>1 \<dots> a\<^sub>n}\<close> prints theorems \<open>a\<^sub>1 \<dots> a\<^sub>n\<close>.
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
219 |
Full fact expressions are allowed here, including attributes |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
220 |
(\secref{sec:syn-att}).
|
| 27043 | 221 |
|
| 61493 | 222 |
\<^descr> \<open>@{prop \<phi>}\<close> prints a well-typed proposition \<open>\<phi>\<close>.
|
| 27043 | 223 |
|
| 61493 | 224 |
\<^descr> \<open>@{lemma \<phi> by m}\<close> proves a well-typed proposition
|
225 |
\<open>\<phi>\<close> by method \<open>m\<close> and prints the original \<open>\<phi>\<close>. |
|
| 27453 | 226 |
|
| 61493 | 227 |
\<^descr> \<open>@{term t}\<close> prints a well-typed term \<open>t\<close>.
|
|
43613
7afbaf5a406b
adding a minimalistic documentation of the value antiquotation in the Isar reference manual
bulwahn
parents:
42936
diff
changeset
|
228 |
|
| 61493 | 229 |
\<^descr> \<open>@{value t}\<close> evaluates a term \<open>t\<close> and prints
|
| 43618 | 230 |
its result, see also @{command_ref (HOL) value}.
|
| 27043 | 231 |
|
| 61493 | 232 |
\<^descr> \<open>@{term_type t}\<close> prints a well-typed term \<open>t\<close>
|
|
32898
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32892
diff
changeset
|
233 |
annotated with its type. |
|
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32892
diff
changeset
|
234 |
|
| 61493 | 235 |
\<^descr> \<open>@{typeof t}\<close> prints the type of a well-typed term
|
236 |
\<open>t\<close>. |
|
|
32898
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32892
diff
changeset
|
237 |
|
| 61493 | 238 |
\<^descr> \<open>@{const c}\<close> prints a logical or syntactic constant
|
239 |
\<open>c\<close>. |
|
| 27043 | 240 |
|
| 61493 | 241 |
\<^descr> \<open>@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}\<close> prints a constant abbreviation
|
242 |
\<open>c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs\<close> as defined in the current context. |
|
|
39305
d4fa19eb0822
'class' and 'type' are now antiquoations by default
haftmann
parents:
32898
diff
changeset
|
243 |
|
| 61493 | 244 |
\<^descr> \<open>@{typ \<tau>}\<close> prints a well-formed type \<open>\<tau>\<close>.
|
|
39305
d4fa19eb0822
'class' and 'type' are now antiquoations by default
haftmann
parents:
32898
diff
changeset
|
245 |
|
| 61493 | 246 |
\<^descr> \<open>@{type \<kappa>}\<close> prints a (logical or syntactic) type
|
247 |
constructor \<open>\<kappa>\<close>. |
|
|
39305
d4fa19eb0822
'class' and 'type' are now antiquoations by default
haftmann
parents:
32898
diff
changeset
|
248 |
|
| 61493 | 249 |
\<^descr> \<open>@{class c}\<close> prints a class \<open>c\<close>.
|
|
39305
d4fa19eb0822
'class' and 'type' are now antiquoations by default
haftmann
parents:
32898
diff
changeset
|
250 |
|
| 61493 | 251 |
\<^descr> \<open>@{goals}\<close> prints the current \<^emph>\<open>dynamic\<close> goal
|
| 27043 | 252 |
state. This is mainly for support of tactic-emulation scripts |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
253 |
within Isar. Presentation of goal states does not conform to the |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
254 |
idea of human-readable proof documents! |
| 27043 | 255 |
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
256 |
When explaining proofs in detail it is usually better to spell out |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
257 |
the reasoning via proper Isar proof commands, instead of peeking at |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
258 |
the internal machine configuration. |
| 27043 | 259 |
|
| 61493 | 260 |
\<^descr> \<open>@{subgoals}\<close> is similar to \<open>@{goals}\<close>, but
|
| 27043 | 261 |
does not print the main goal. |
262 |
||
| 61493 | 263 |
\<^descr> \<open>@{prf a\<^sub>1 \<dots> a\<^sub>n}\<close> prints the (compact) proof terms
|
264 |
corresponding to the theorems \<open>a\<^sub>1 \<dots> a\<^sub>n\<close>. Note that this |
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
265 |
requires proof terms to be switched on for the current logic |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
266 |
session. |
| 27043 | 267 |
|
| 61493 | 268 |
\<^descr> \<open>@{full_prf a\<^sub>1 \<dots> a\<^sub>n}\<close> is like \<open>@{prf a\<^sub>1 \<dots>
|
269 |
a\<^sub>n}\<close>, but prints the full proof terms, i.e.\ also displays |
|
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
270 |
information omitted in the compact proof term, which is denoted by |
| 61493 | 271 |
``\<open>_\<close>'' placeholders there. |
| 27043 | 272 |
|
| 61493 | 273 |
\<^descr> \<open>@{ML s}\<close>, \<open>@{ML_op s}\<close>, \<open>@{ML_type
|
274 |
s}\<close>, \<open>@{ML_structure s}\<close>, and \<open>@{ML_functor s}\<close>
|
|
275 |
check text \<open>s\<close> as ML value, infix operator, type, structure, |
|
| 55837 | 276 |
and functor respectively. The source is printed verbatim. |
| 27043 | 277 |
|
| 61493 | 278 |
\<^descr> \<open>@{emph s}\<close> prints document source recursively, with {\LaTeX}
|
| 61503 | 279 |
markup \<^verbatim>\<open>\emph{\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close>.
|
|
61473
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
280 |
|
| 61493 | 281 |
\<^descr> \<open>@{bold s}\<close> prints document source recursively, with {\LaTeX}
|
| 61503 | 282 |
markup \<^verbatim>\<open>\textbf{\<close>\<open>\<dots>\<close>\<^verbatim>\<open>}\<close>.
|
|
61473
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
283 |
|
| 61493 | 284 |
\<^descr> \<open>@{verbatim s}\<close> prints uninterpreted source text literally
|
|
58716
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents:
58618
diff
changeset
|
285 |
as ASCII characters, using some type-writer font style. |
|
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
wenzelm
parents:
58618
diff
changeset
|
286 |
|
| 61493 | 287 |
\<^descr> \<open>@{file path}\<close> checks that \<open>path\<close> refers to a
|
| 40801 | 288 |
file (or directory) and prints it verbatim. |
289 |
||
| 61493 | 290 |
\<^descr> \<open>@{file_unchecked path}\<close> is like \<open>@{file
|
291 |
path}\<close>, but does not check the existence of the \<open>path\<close> |
|
| 54705 | 292 |
within the file-system. |
293 |
||
| 61493 | 294 |
\<^descr> \<open>@{url name}\<close> produces markup for the given URL, which
|
|
54702
3daeba5130f0
added document antiquotation @{url}, which produces formal markup for LaTeX and PIDE;
wenzelm
parents:
54346
diff
changeset
|
295 |
results in an active hyperlink within the text. |
|
3daeba5130f0
added document antiquotation @{url}, which produces formal markup for LaTeX and PIDE;
wenzelm
parents:
54346
diff
changeset
|
296 |
|
| 61503 | 297 |
\<^descr> \<open>@{cite name}\<close> produces a citation \<^verbatim>\<open>\cite{name}\<close> in {\LaTeX},
|
298 |
where the name refers to some Bib{\TeX} database entry.
|
|
| 58593 | 299 |
|
| 61503 | 300 |
The variant \<open>@{cite \<open>opt\<close> name}\<close> produces \<^verbatim>\<open>\cite[opt]{name}\<close> with
|
301 |
some free-form optional argument. Multiple names |
|
| 61493 | 302 |
are output with commas, e.g. \<open>@{cite foo \<AND> bar}\<close> becomes
|
| 61503 | 303 |
\<^verbatim>\<open>\cite{foo,bar}\<close>.
|
| 58593 | 304 |
|
305 |
The {\LaTeX} macro name is determined by the antiquotation option
|
|
306 |
@{antiquotation_option_def cite_macro}, or the configuration option
|
|
| 61493 | 307 |
@{attribute cite_macro} in the context. For example, \<open>@{cite
|
| 61503 | 308 |
[cite_macro = nocite] foobar}\<close> produces \<^verbatim>\<open>\nocite{foobar}\<close>.
|
|
61473
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
309 |
|
|
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
310 |
\<^descr> @{command "print_antiquotations"} prints all document antiquotations
|
| 61493 | 311 |
that are defined in the current context; the ``\<open>!\<close>'' option |
|
61473
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
wenzelm
parents:
61472
diff
changeset
|
312 |
indicates extra verbosity. |
| 58618 | 313 |
\<close> |
| 27043 | 314 |
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
315 |
|
| 58618 | 316 |
subsection \<open>Styled antiquotations\<close> |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
317 |
|
| 61493 | 318 |
text \<open>The antiquotations \<open>thm\<close>, \<open>prop\<close> and \<open>term\<close> admit an extra \<^emph>\<open>style\<close> specification to modify the |
| 32891 | 319 |
printed result. A style is specified by a name with a possibly |
320 |
empty number of arguments; multiple styles can be sequenced with |
|
321 |
commas. The following standard styles are available: |
|
| 27043 | 322 |
|
| 61493 | 323 |
\<^descr> \<open>lhs\<close> extracts the first argument of any application |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
324 |
form with at least two arguments --- typically meta-level or |
| 27043 | 325 |
object-level equality, or any other binary relation. |
326 |
||
| 61493 | 327 |
\<^descr> \<open>rhs\<close> is like \<open>lhs\<close>, but extracts the second |
| 27043 | 328 |
argument. |
329 |
||
| 61493 | 330 |
\<^descr> \<open>concl\<close> extracts the conclusion \<open>C\<close> from a rule |
331 |
in Horn-clause normal form \<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C\<close>. |
|
| 27043 | 332 |
|
| 61493 | 333 |
\<^descr> \<open>prem\<close> \<open>n\<close> extract premise number |
334 |
\<open>n\<close> from from a rule in Horn-clause |
|
335 |
normal form \<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C\<close> |
|
| 58618 | 336 |
\<close> |
| 27043 | 337 |
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
338 |
|
| 58618 | 339 |
subsection \<open>General options\<close> |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
340 |
|
| 58618 | 341 |
text \<open>The following options are available to tune the printed output |
| 54346 | 342 |
of antiquotations. Note that many of these coincide with system and |
343 |
configuration options of the same names. |
|
| 27043 | 344 |
|
| 61493 | 345 |
\<^descr> @{antiquotation_option_def show_types}~\<open>= bool\<close> and
|
346 |
@{antiquotation_option_def show_sorts}~\<open>= bool\<close> control
|
|
| 30397 | 347 |
printing of explicit type and sort constraints. |
| 27043 | 348 |
|
| 61493 | 349 |
\<^descr> @{antiquotation_option_def show_structs}~\<open>= bool\<close>
|
| 30397 | 350 |
controls printing of implicit structures. |
| 27043 | 351 |
|
| 61493 | 352 |
\<^descr> @{antiquotation_option_def show_abbrevs}~\<open>= bool\<close>
|
|
40879
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents:
40801
diff
changeset
|
353 |
controls folding of abbreviations. |
|
ca132ef44944
configuration option "show_abbrevs" supersedes print mode "no_abbrevs", with inverted meaning;
wenzelm
parents:
40801
diff
changeset
|
354 |
|
| 61493 | 355 |
\<^descr> @{antiquotation_option_def names_long}~\<open>= bool\<close> forces
|
| 30397 | 356 |
names of types and constants etc.\ to be printed in their fully |
357 |
qualified internal form. |
|
| 27043 | 358 |
|
| 61493 | 359 |
\<^descr> @{antiquotation_option_def names_short}~\<open>= bool\<close>
|
| 30397 | 360 |
forces names of types and constants etc.\ to be printed unqualified. |
361 |
Note that internalizing the output again in the current context may |
|
362 |
well yield a different result. |
|
| 27043 | 363 |
|
| 61493 | 364 |
\<^descr> @{antiquotation_option_def names_unique}~\<open>= bool\<close>
|
| 30397 | 365 |
determines whether the printed version of qualified names should be |
366 |
made sufficiently long to avoid overlap with names declared further |
|
| 61493 | 367 |
back. Set to \<open>false\<close> for more concise output. |
| 27043 | 368 |
|
| 61493 | 369 |
\<^descr> @{antiquotation_option_def eta_contract}~\<open>= bool\<close>
|
370 |
prints terms in \<open>\<eta>\<close>-contracted form. |
|
| 27043 | 371 |
|
| 61493 | 372 |
\<^descr> @{antiquotation_option_def display}~\<open>= bool\<close> indicates
|
| 30397 | 373 |
if the text is to be output as multi-line ``display material'', |
374 |
rather than a small piece of text without line breaks (which is the |
|
375 |
default). |
|
| 27043 | 376 |
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
377 |
In this mode the embedded entities are printed in the same style as |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
378 |
the main theory text. |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
379 |
|
| 61493 | 380 |
\<^descr> @{antiquotation_option_def break}~\<open>= bool\<close> controls
|
| 30397 | 381 |
line breaks in non-display material. |
| 27043 | 382 |
|
| 61493 | 383 |
\<^descr> @{antiquotation_option_def quotes}~\<open>= bool\<close> indicates
|
| 30397 | 384 |
if the output should be enclosed in double quotes. |
| 27043 | 385 |
|
| 61493 | 386 |
\<^descr> @{antiquotation_option_def mode}~\<open>= name\<close> adds \<open>name\<close> to the print mode to be used for presentation. Note that the
|
| 30397 | 387 |
standard setup for {\LaTeX} output is already present by default,
|
| 61493 | 388 |
including the modes \<open>latex\<close> and \<open>xsymbols\<close>. |
| 27043 | 389 |
|
| 61493 | 390 |
\<^descr> @{antiquotation_option_def margin}~\<open>= nat\<close> and
|
391 |
@{antiquotation_option_def indent}~\<open>= nat\<close> change the margin
|
|
| 30397 | 392 |
or indentation for pretty printing of display material. |
| 27043 | 393 |
|
| 61493 | 394 |
\<^descr> @{antiquotation_option_def goals_limit}~\<open>= nat\<close>
|
|
51960
61ac1efe02c3
option "goals_limit", with more uniform description;
wenzelm
parents:
51057
diff
changeset
|
395 |
determines the maximum number of subgoals to be printed (for goal-based |
| 30397 | 396 |
antiquotation). |
| 27043 | 397 |
|
| 61493 | 398 |
\<^descr> @{antiquotation_option_def source}~\<open>= bool\<close> prints the
|
| 30397 | 399 |
original source text of the antiquotation arguments, rather than its |
400 |
internal representation. Note that formal checking of |
|
401 |
@{antiquotation "thm"}, @{antiquotation "term"}, etc. is still
|
|
402 |
enabled; use the @{antiquotation "text"} antiquotation for unchecked
|
|
403 |
output. |
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
404 |
|
| 61493 | 405 |
Regular \<open>term\<close> and \<open>typ\<close> antiquotations with \<open>source = false\<close> involve a full round-trip from the original source |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
406 |
to an internalized logical entity back to a source form, according |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
407 |
to the syntax of the current context. Thus the printed output is |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
408 |
not under direct control of the author, it may even fluctuate a bit |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
409 |
as the underlying theory is changed later on. |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
410 |
|
| 61493 | 411 |
In contrast, @{antiquotation_option source}~\<open>= true\<close>
|
| 30397 | 412 |
admits direct printing of the given source text, with the desirable |
413 |
well-formedness check in the background, but without modification of |
|
414 |
the printed text. |
|
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
415 |
|
| 27043 | 416 |
|
| 61493 | 417 |
For Boolean flags, ``\<open>name = true\<close>'' may be abbreviated as |
418 |
``\<open>name\<close>''. All of the above flags are disabled by default, |
|
| 51057 | 419 |
unless changed specifically for a logic session in the corresponding |
| 61503 | 420 |
\<^verbatim>\<open>ROOT\<close> file. |
| 61458 | 421 |
\<close> |
| 27043 | 422 |
|
423 |
||
| 58618 | 424 |
section \<open>Markup via command tags \label{sec:tags}\<close>
|
| 27043 | 425 |
|
| 58618 | 426 |
text \<open>Each Isabelle/Isar command may be decorated by additional |
| 28750 | 427 |
presentation tags, to indicate some modification in the way it is |
428 |
printed in the document. |
|
| 27043 | 429 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
430 |
@{rail \<open>
|
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
431 |
@{syntax_def tags}: ( tag * )
|
| 27043 | 432 |
; |
|
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40879
diff
changeset
|
433 |
tag: '%' (@{syntax ident} | @{syntax string})
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
434 |
\<close>} |
| 27043 | 435 |
|
| 28750 | 436 |
Some tags are pre-declared for certain classes of commands, serving |
437 |
as default markup if no tags are given in the text: |
|
| 27043 | 438 |
|
| 61421 | 439 |
\<^medskip> |
| 27043 | 440 |
\begin{tabular}{ll}
|
| 61493 | 441 |
\<open>theory\<close> & theory begin/end \\ |
442 |
\<open>proof\<close> & all proof commands \\ |
|
443 |
\<open>ML\<close> & all commands involving ML code \\ |
|
| 27043 | 444 |
\end{tabular}
|
| 61421 | 445 |
\<^medskip> |
| 27043 | 446 |
|
| 61421 | 447 |
The Isabelle document preparation system |
| 60270 | 448 |
@{cite "isabelle-system"} allows tagged command regions to be presented
|
| 27043 | 449 |
specifically, e.g.\ to fold proof texts, or drop parts of the text |
450 |
completely. |
|
451 |
||
| 61493 | 452 |
For example ``@{command "by"}~\<open>%invisible auto\<close>'' causes
|
453 |
that piece of proof to be treated as \<open>invisible\<close> instead of |
|
454 |
\<open>proof\<close> (the default), which may be shown or hidden |
|
| 27043 | 455 |
depending on the document setup. In contrast, ``@{command
|
| 61493 | 456 |
"by"}~\<open>%visible auto\<close>'' forces this text to be shown |
| 27043 | 457 |
invariably. |
458 |
||
459 |
Explicit tag specifications within a proof apply to all subsequent |
|
460 |
commands of the same level of nesting. For example, ``@{command
|
|
| 61493 | 461 |
"proof"}~\<open>%visible \<dots>\<close>~@{command "qed"}'' forces the whole
|
462 |
sub-proof to be typeset as \<open>visible\<close> (unless some of its parts |
|
| 28750 | 463 |
are tagged differently). |
464 |
||
| 61421 | 465 |
\<^medskip> |
466 |
Command tags merely produce certain markup environments for |
|
| 28750 | 467 |
type-setting. The meaning of these is determined by {\LaTeX}
|
|
40800
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
wenzelm
parents:
40255
diff
changeset
|
468 |
macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or
|
| 28750 | 469 |
by the document author. The Isabelle document preparation tools |
470 |
also provide some high-level options to specify the meaning of |
|
471 |
arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding |
|
472 |
parts of the text. Logic sessions may also specify ``document |
|
473 |
versions'', where given tags are interpreted in some particular way. |
|
| 60270 | 474 |
Again see @{cite "isabelle-system"} for further details.
|
| 58618 | 475 |
\<close> |
| 27043 | 476 |
|
477 |
||
| 58618 | 478 |
section \<open>Railroad diagrams\<close> |
| 42658 | 479 |
|
| 58618 | 480 |
text \<open> |
| 42658 | 481 |
\begin{matharray}{rcl}
|
| 61493 | 482 |
@{antiquotation_def "rail"} & : & \<open>antiquotation\<close> \\
|
| 42658 | 483 |
\end{matharray}
|
484 |
||
| 55113 | 485 |
@{rail \<open>
|
| 59937 | 486 |
'rail' @{syntax text}
|
| 55113 | 487 |
\<close>} |
| 42658 | 488 |
|
489 |
The @{antiquotation rail} antiquotation allows to include syntax
|
|
490 |
diagrams into Isabelle documents. {\LaTeX} requires the style file
|
|
| 59116 | 491 |
@{file "~~/lib/texinputs/railsetup.sty"}, which can be used via
|
| 61503 | 492 |
\<^verbatim>\<open>\usepackage{railsetup}\<close> in \<^verbatim>\<open>root.tex\<close>, for example.
|
| 42658 | 493 |
|
494 |
The rail specification language is quoted here as Isabelle @{syntax
|
|
| 55120 | 495 |
string} or text @{syntax "cartouche"}; it has its own grammar given
|
496 |
below. |
|
| 42658 | 497 |
|
| 55120 | 498 |
\begingroup |
| 61474 | 499 |
\def\isasymnewline{\isatt{\isacharbackslash\isacharless newline\isachargreater}}
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
500 |
@{rail \<open>
|
| 42658 | 501 |
rule? + ';' |
502 |
; |
|
503 |
rule: ((identifier | @{syntax antiquotation}) ':')? body
|
|
504 |
; |
|
505 |
body: concatenation + '|' |
|
506 |
; |
|
507 |
concatenation: ((atom '?'?) +) (('*' | '+') atom?)?
|
|
508 |
; |
|
509 |
atom: '(' body? ')' | identifier |
|
|
510 |
'@'? (string | @{syntax antiquotation}) |
|
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
511 |
'\<newline>' |
|
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
512 |
\<close>} |
| 55120 | 513 |
\endgroup |
| 42658 | 514 |
|
| 61493 | 515 |
The lexical syntax of \<open>identifier\<close> coincides with that of |
516 |
@{syntax ident} in regular Isabelle syntax, but \<open>string\<close> uses
|
|
| 42658 | 517 |
single quotes instead of double quotes of the standard @{syntax
|
| 55113 | 518 |
string} category. |
| 42658 | 519 |
|
| 61493 | 520 |
Each \<open>rule\<close> defines a formal language (with optional name), |
| 42658 | 521 |
using a notation that is similar to EBNF or regular expressions with |
522 |
recursion. The meaning and visual appearance of these rail language |
|
523 |
elements is illustrated by the following representative examples. |
|
524 |
||
| 61503 | 525 |
\<^item> Empty \<^verbatim>\<open>()\<close> |
| 42658 | 526 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
527 |
@{rail \<open>()\<close>}
|
| 42658 | 528 |
|
| 61503 | 529 |
\<^item> Nonterminal \<^verbatim>\<open>A\<close> |
| 42658 | 530 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
531 |
@{rail \<open>A\<close>}
|
| 42658 | 532 |
|
| 61503 | 533 |
\<^item> Nonterminal via Isabelle antiquotation \<^verbatim>\<open>@{syntax method}\<close>
|
| 42658 | 534 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
535 |
@{rail \<open>@{syntax method}\<close>}
|
| 42658 | 536 |
|
| 61503 | 537 |
\<^item> Terminal \<^verbatim>\<open>'xyz'\<close> |
| 42658 | 538 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
539 |
@{rail \<open>'xyz'\<close>}
|
| 42658 | 540 |
|
| 61503 | 541 |
\<^item> Terminal in keyword style \<^verbatim>\<open>@'xyz'\<close> |
| 42658 | 542 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
543 |
@{rail \<open>@'xyz'\<close>}
|
| 42658 | 544 |
|
| 61503 | 545 |
\<^item> Terminal via Isabelle antiquotation \<^verbatim>\<open>@@{method rule}\<close>
|
| 42658 | 546 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
547 |
@{rail \<open>@@{method rule}\<close>}
|
| 42658 | 548 |
|
| 61503 | 549 |
\<^item> Concatenation \<^verbatim>\<open>A B C\<close> |
| 42658 | 550 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
551 |
@{rail \<open>A B C\<close>}
|
| 42658 | 552 |
|
| 61503 | 553 |
\<^item> Newline inside concatenation \<^verbatim>\<open>A B C \<newline> D E F\<close> |
| 42658 | 554 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
555 |
@{rail \<open>A B C \<newline> D E F\<close>}
|
| 42658 | 556 |
|
| 61503 | 557 |
\<^item> Variants \<^verbatim>\<open>A | B | C\<close> |
| 42658 | 558 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
559 |
@{rail \<open>A | B | C\<close>}
|
| 42658 | 560 |
|
| 61503 | 561 |
\<^item> Option \<^verbatim>\<open>A ?\<close> |
| 42658 | 562 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
563 |
@{rail \<open>A ?\<close>}
|
| 42658 | 564 |
|
| 61503 | 565 |
\<^item> Repetition \<^verbatim>\<open>A *\<close> |
| 42658 | 566 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
567 |
@{rail \<open>A *\<close>}
|
| 42658 | 568 |
|
| 61503 | 569 |
\<^item> Repetition with separator \<^verbatim>\<open>A * sep\<close> |
| 42658 | 570 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
571 |
@{rail \<open>A * sep\<close>}
|
| 42658 | 572 |
|
| 61503 | 573 |
\<^item> Strict repetition \<^verbatim>\<open>A +\<close> |
| 42658 | 574 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
575 |
@{rail \<open>A +\<close>}
|
| 42658 | 576 |
|
| 61503 | 577 |
\<^item> Strict repetition with separator \<^verbatim>\<open>A + sep\<close> |
| 42658 | 578 |
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
579 |
@{rail \<open>A + sep\<close>}
|
| 58618 | 580 |
\<close> |
| 42658 | 581 |
|
582 |
||
| 58618 | 583 |
section \<open>Draft presentation\<close> |
| 27043 | 584 |
|
| 58618 | 585 |
text \<open> |
| 27043 | 586 |
\begin{matharray}{rcl}
|
| 61493 | 587 |
@{command_def "display_drafts"}\<open>\<^sup>*\<close> & : & \<open>any \<rightarrow>\<close> \\
|
| 27043 | 588 |
\end{matharray}
|
589 |
||
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
590 |
@{rail \<open>
|
| 52549 | 591 |
@@{command display_drafts} (@{syntax name} +)
|
|
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
592 |
\<close>} |
| 27043 | 593 |
|
| 61493 | 594 |
\<^descr> @{command "display_drafts"}~\<open>paths\<close> performs simple output of a
|
| 52549 | 595 |
given list of raw source files. Only those symbols that do not require |
596 |
additional {\LaTeX} packages are displayed properly, everything else is left
|
|
597 |
verbatim. |
|
| 58618 | 598 |
\<close> |
| 27043 | 599 |
|
600 |
end |