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