author | wenzelm |
Fri, 13 Mar 2009 23:56:07 +0100 | |
changeset 30521 | 3ec2d35b380f |
parent 30397 | b6212ae21656 |
child 32891 | d403b99287ff |
permissions | -rw-r--r-- |
27043 | 1 |
theory Document_Preparation |
2 |
imports Main |
|
3 |
begin |
|
4 |
||
5 |
chapter {* Document preparation \label{ch:document-prep} *} |
|
6 |
||
28746 | 7 |
text {* Isabelle/Isar provides a simple document preparation system |
8 |
based on regular {PDF-\LaTeX} technology, with full support for |
|
9 |
hyper-links and bookmarks. Thus the results are well suited for WWW |
|
10 |
browsing and as printed copies. |
|
27043 | 11 |
|
28746 | 12 |
\medskip Isabelle generates {\LaTeX} output while running a |
27043 | 13 |
\emph{logic session} (see also \cite{isabelle-sys}). Getting |
14 |
started with a working configuration for common situations is quite |
|
15 |
easy by using the Isabelle @{verbatim mkdir} and @{verbatim make} |
|
16 |
tools. First invoke |
|
17 |
\begin{ttbox} |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
27881
diff
changeset
|
18 |
isabelle mkdir Foo |
27043 | 19 |
\end{ttbox} |
28746 | 20 |
to initialize a separate directory for session @{verbatim Foo} (it |
21 |
is safe to experiment, since @{verbatim "isabelle mkdir"} never |
|
22 |
overwrites existing files). Ensure that @{verbatim "Foo/ROOT.ML"} |
|
27043 | 23 |
holds ML commands to load all theories required for this session; |
24 |
furthermore @{verbatim "Foo/document/root.tex"} should include any |
|
25 |
special {\LaTeX} macro packages required for your document (the |
|
26 |
default is usually sufficient as a start). |
|
27 |
||
28 |
The session is controlled by a separate @{verbatim IsaMakefile} |
|
29 |
(with crude source dependencies by default). This file is located |
|
30 |
one level up from the @{verbatim Foo} directory location. Now |
|
31 |
invoke |
|
32 |
\begin{ttbox} |
|
28504
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
wenzelm
parents:
27881
diff
changeset
|
33 |
isabelle make Foo |
27043 | 34 |
\end{ttbox} |
35 |
to run the @{verbatim Foo} session, with browser information and |
|
36 |
document preparation enabled. Unless any errors are reported by |
|
37 |
Isabelle or {\LaTeX}, the output will appear inside the directory |
|
28746 | 38 |
defined by the @{verbatim ISABELLE_BROWSER_INFO} setting (as |
39 |
reported by the batch job in verbose mode). |
|
27043 | 40 |
|
41 |
\medskip You may also consider to tune the @{verbatim usedir} |
|
28746 | 42 |
options in @{verbatim IsaMakefile}, for example to switch the output |
43 |
format between @{verbatim pdf} and @{verbatim dvi}, or activate the |
|
27043 | 44 |
@{verbatim "-D"} option to retain a second copy of the generated |
28746 | 45 |
{\LaTeX} sources (for manual inspection or separate runs of |
46 |
@{executable latex}). |
|
27043 | 47 |
|
48 |
\medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys} |
|
49 |
for further details on Isabelle logic sessions and theory |
|
50 |
presentation. The Isabelle/HOL tutorial \cite{isabelle-hol-book} |
|
28746 | 51 |
also covers theory presentation to some extent. |
27043 | 52 |
*} |
53 |
||
54 |
||
55 |
section {* Markup commands \label{sec:markup} *} |
|
56 |
||
57 |
text {* |
|
58 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
59 |
@{command_def "header"} & : & @{text "toplevel \<rightarrow> toplevel"} \\[0.5ex] |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
60 |
@{command_def "chapter"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
61 |
@{command_def "section"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
62 |
@{command_def "subsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
63 |
@{command_def "subsubsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
64 |
@{command_def "text"} & : & @{text "local_theory \<rightarrow> local_theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
65 |
@{command_def "text_raw"} & : & @{text "local_theory \<rightarrow> local_theory"} \\[0.5ex] |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
66 |
@{command_def "sect"} & : & @{text "proof \<rightarrow> proof"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
67 |
@{command_def "subsect"} & : & @{text "proof \<rightarrow> proof"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
68 |
@{command_def "subsubsect"} & : & @{text "proof \<rightarrow> proof"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
69 |
@{command_def "txt"} & : & @{text "proof \<rightarrow> proof"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
70 |
@{command_def "txt_raw"} & : & @{text "proof \<rightarrow> proof"} \\ |
27043 | 71 |
\end{matharray} |
72 |
||
28747 | 73 |
Markup commands provide a structured way to insert text into the |
74 |
document generated from a theory. Each markup command takes a |
|
75 |
single @{syntax text} argument, which is passed as argument to a |
|
76 |
corresponding {\LaTeX} macro. The default macros provided by |
|
77 |
@{"file" "~~/lib/texinputs/isabelle.sty"} can be redefined according |
|
78 |
to the needs of the underlying document and {\LaTeX} styles. |
|
79 |
||
80 |
Note that formal comments (\secref{sec:comments}) are similar to |
|
81 |
markup commands, but have a different status within Isabelle/Isar |
|
82 |
syntax. |
|
27043 | 83 |
|
84 |
\begin{rail} |
|
85 |
('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text') target? text |
|
86 |
; |
|
27049 | 87 |
('header' | 'text\_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt\_raw') text |
27043 | 88 |
; |
89 |
\end{rail} |
|
90 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
91 |
\begin{description} |
27043 | 92 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
93 |
\item @{command header} provides plain text markup just preceding |
28747 | 94 |
the formal beginning of a theory. The corresponding {\LaTeX} macro |
95 |
is @{verbatim "\\isamarkupheader"}, which acts like @{command |
|
96 |
section} by default. |
|
27049 | 97 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
98 |
\item @{command chapter}, @{command section}, @{command subsection}, |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
99 |
and @{command subsubsection} mark chapter and section headings |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
100 |
within the main theory body or local theory targets. The |
28747 | 101 |
corresponding {\LaTeX} macros are @{verbatim "\\isamarkupchapter"}, |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
102 |
@{verbatim "\\isamarkupsection"}, @{verbatim |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
103 |
"\\isamarkupsubsection"} etc. |
28747 | 104 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
105 |
\item @{command sect}, @{command subsect}, and @{command subsubsect} |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
106 |
mark section headings within proofs. The corresponding {\LaTeX} |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
107 |
macros are @{verbatim "\\isamarkupsect"}, @{verbatim |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
108 |
"\\isamarkupsubsect"} etc. |
27043 | 109 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
110 |
\item @{command text} and @{command txt} specify paragraphs of plain |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
111 |
text. This corresponds to a {\LaTeX} environment @{verbatim |
28747 | 112 |
"\\begin{isamarkuptext}"} @{text "\<dots>"} @{verbatim |
113 |
"\\end{isamarkuptext}"} etc. |
|
27043 | 114 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
115 |
\item @{command text_raw} and @{command txt_raw} insert {\LaTeX} |
28747 | 116 |
source into the output, without additional markup. Thus the full |
117 |
range of document manipulations becomes available, at the risk of |
|
118 |
messing up document output. |
|
27043 | 119 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
120 |
\end{description} |
27043 | 121 |
|
28747 | 122 |
Except for @{command "text_raw"} and @{command "txt_raw"}, the text |
123 |
passed to any of the above markup commands may refer to formal |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
124 |
entities via \emph{document antiquotations}, see also |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
125 |
\secref{sec:antiq}. These are interpreted in the present theory or |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
126 |
proof context, or the named @{text "target"}. |
27043 | 127 |
|
128 |
\medskip The proof markup commands closely resemble those for theory |
|
129 |
specifications, but have a different formal status and produce |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
130 |
different {\LaTeX} macros. The default definitions coincide for |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
131 |
analogous commands such as @{command section} and @{command sect}. |
27043 | 132 |
*} |
133 |
||
134 |
||
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
135 |
section {* Document Antiquotations \label{sec:antiq} *} |
27043 | 136 |
|
137 |
text {* |
|
138 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
139 |
@{antiquotation_def "theory"} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
140 |
@{antiquotation_def "thm"} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
141 |
@{antiquotation_def "lemma"} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
142 |
@{antiquotation_def "prop"} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
143 |
@{antiquotation_def "term"} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
144 |
@{antiquotation_def const} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
145 |
@{antiquotation_def abbrev} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
146 |
@{antiquotation_def typeof} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
147 |
@{antiquotation_def typ} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
148 |
@{antiquotation_def thm_style} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
149 |
@{antiquotation_def term_style} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
150 |
@{antiquotation_def "text"} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
151 |
@{antiquotation_def goals} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
152 |
@{antiquotation_def subgoals} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
153 |
@{antiquotation_def prf} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
154 |
@{antiquotation_def full_prf} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
155 |
@{antiquotation_def ML} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
156 |
@{antiquotation_def ML_type} & : & @{text antiquotation} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
157 |
@{antiquotation_def ML_struct} & : & @{text antiquotation} \\ |
27043 | 158 |
\end{matharray} |
159 |
||
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
160 |
The overall content of an Isabelle/Isar theory may alternate between |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
161 |
formal and informal text. The main body consists of formal |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
162 |
specification and proof commands, interspersed with markup commands |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
163 |
(\secref{sec:markup}) or document comments (\secref{sec:comments}). |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
164 |
The argument of markup commands quotes informal text to be printed |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
165 |
in the resulting document, but may again refer to formal entities |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
166 |
via \emph{document antiquotations}. |
27043 | 167 |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
168 |
For example, embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}'' |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
169 |
within a text block makes |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
170 |
\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
|
171 |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
172 |
Antiquotations usually spare the author tedious typing of logical |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
173 |
entities in full detail. Even more importantly, some degree of |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
174 |
consistency-checking between the main body of formal text and its |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
175 |
informal explanation is achieved, since terms and types appearing in |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
176 |
antiquotations are checked within the current theory or proof |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
177 |
context. |
27043 | 178 |
|
179 |
\begin{rail} |
|
180 |
atsign lbrace antiquotation rbrace |
|
181 |
; |
|
182 |
||
183 |
antiquotation: |
|
184 |
'theory' options name | |
|
185 |
'thm' options thmrefs | |
|
27453 | 186 |
'lemma' options prop 'by' method | |
27043 | 187 |
'prop' options prop | |
188 |
'term' options term | |
|
189 |
'const' options term | |
|
190 |
'abbrev' options term | |
|
191 |
'typeof' options term | |
|
192 |
'typ' options type | |
|
193 |
'thm\_style' options name thmref | |
|
194 |
'term\_style' options name term | |
|
195 |
'text' options name | |
|
196 |
'goals' options | |
|
197 |
'subgoals' options | |
|
198 |
'prf' options thmrefs | |
|
199 |
'full\_prf' options thmrefs | |
|
200 |
'ML' options name | |
|
201 |
'ML\_type' options name | |
|
202 |
'ML\_struct' options name |
|
203 |
; |
|
204 |
options: '[' (option * ',') ']' |
|
205 |
; |
|
206 |
option: name | name '=' name |
|
207 |
; |
|
208 |
\end{rail} |
|
209 |
||
210 |
Note that the syntax of antiquotations may \emph{not} include source |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
211 |
comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} nor verbatim |
27043 | 212 |
text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim |
213 |
"*"}@{verbatim "}"}. |
|
214 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
215 |
\begin{description} |
27043 | 216 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
217 |
\item @{text "@{theory A}"} prints the name @{text "A"}, which is |
27043 | 218 |
guaranteed to refer to a valid ancestor theory in the current |
219 |
context. |
|
220 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
221 |
\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
|
222 |
Full fact expressions are allowed here, including attributes |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
223 |
(\secref{sec:syn-att}). |
27043 | 224 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
225 |
\item @{text "@{prop \<phi>}"} prints a well-typed proposition @{text |
27043 | 226 |
"\<phi>"}. |
227 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
228 |
\item @{text "@{lemma \<phi> by m}"} proves a well-typed proposition |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
229 |
@{text "\<phi>"} by method @{text m} and prints the original @{text "\<phi>"}. |
27453 | 230 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
231 |
\item @{text "@{term t}"} prints a well-typed term @{text "t"}. |
27043 | 232 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
233 |
\item @{text "@{const c}"} prints a logical or syntactic constant |
27043 | 234 |
@{text "c"}. |
235 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
236 |
\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
|
237 |
@{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in the current context. |
27043 | 238 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
239 |
\item @{text "@{typeof t}"} prints the type of a well-typed term |
27043 | 240 |
@{text "t"}. |
241 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
242 |
\item @{text "@{typ \<tau>}"} prints a well-formed type @{text "\<tau>"}. |
27043 | 243 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
244 |
\item @{text "@{thm_style s a}"} prints theorem @{text a}, |
27043 | 245 |
previously applying a style @{text s} to it (see below). |
246 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
247 |
\item @{text "@{term_style s t}"} prints a well-typed term @{text t} |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
248 |
after applying a style @{text s} to it (see below). |
27043 | 249 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
250 |
\item @{text "@{text s}"} prints uninterpreted source text @{text |
27043 | 251 |
s}. This is particularly useful to print portions of text according |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
252 |
to the Isabelle document style, without demanding well-formedness, |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
253 |
e.g.\ small pieces of terms that should not be parsed or |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
254 |
type-checked yet. |
27043 | 255 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
256 |
\item @{text "@{goals}"} prints the current \emph{dynamic} goal |
27043 | 257 |
state. This is mainly for support of tactic-emulation scripts |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
258 |
within Isar. Presentation of goal states does not conform to the |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
259 |
idea of human-readable proof documents! |
27043 | 260 |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
261 |
When explaining proofs in detail it is usually better to spell out |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
262 |
the reasoning via proper Isar proof commands, instead of peeking at |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
263 |
the internal machine configuration. |
27043 | 264 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
265 |
\item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but |
27043 | 266 |
does not print the main goal. |
267 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
268 |
\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
|
269 |
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
|
270 |
requires proof terms to be switched on for the current logic |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
271 |
session. |
27043 | 272 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
273 |
\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
|
274 |
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
|
275 |
information omitted in the compact proof term, which is denoted by |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
276 |
``@{text _}'' placeholders there. |
27043 | 277 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
278 |
\item @{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
279 |
"@{ML_struct s}"} check text @{text s} as ML value, type, and |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
280 |
structure, respectively. The source is printed verbatim. |
27043 | 281 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
282 |
\end{description} |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
283 |
*} |
27043 | 284 |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
285 |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
286 |
subsubsection {* Styled antiquotations *} |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
287 |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
288 |
text {* Some antiquotations like @{text thm_style} and @{text |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
289 |
term_style} admit an extra \emph{style} specification to modify the |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
290 |
printed result. The following standard styles are available: |
27043 | 291 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
292 |
\begin{description} |
27043 | 293 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
294 |
\item @{text lhs} extracts the first argument of any application |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
295 |
form with at least two arguments --- typically meta-level or |
27043 | 296 |
object-level equality, or any other binary relation. |
297 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
298 |
\item @{text rhs} is like @{text lhs}, but extracts the second |
27043 | 299 |
argument. |
300 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
301 |
\item @{text "concl"} extracts the conclusion @{text C} from a rule |
27043 | 302 |
in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}. |
303 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
304 |
\item @{text "prem1"}, \dots, @{text "prem9"} extract premise number |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
305 |
@{text "1, \<dots>, 9"}, respectively, from from a rule in Horn-clause |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
306 |
normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"} |
27043 | 307 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
308 |
\end{description} |
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
309 |
*} |
27043 | 310 |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
311 |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
312 |
subsubsection {* General options *} |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
313 |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
314 |
text {* The following options are available to tune the printed output |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
315 |
of antiquotations. Note that many of these coincide with global ML |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
316 |
flags of the same names. |
27043 | 317 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
318 |
\begin{description} |
27043 | 319 |
|
30397 | 320 |
\item @{antiquotation_option_def show_types}~@{text "= bool"} and |
321 |
@{antiquotation_option_def show_sorts}~@{text "= bool"} control |
|
322 |
printing of explicit type and sort constraints. |
|
27043 | 323 |
|
30397 | 324 |
\item @{antiquotation_option_def show_structs}~@{text "= bool"} |
325 |
controls printing of implicit structures. |
|
27043 | 326 |
|
30397 | 327 |
\item @{antiquotation_option_def long_names}~@{text "= bool"} forces |
328 |
names of types and constants etc.\ to be printed in their fully |
|
329 |
qualified internal form. |
|
27043 | 330 |
|
30397 | 331 |
\item @{antiquotation_option_def short_names}~@{text "= bool"} |
332 |
forces names of types and constants etc.\ to be printed unqualified. |
|
333 |
Note that internalizing the output again in the current context may |
|
334 |
well yield a different result. |
|
27043 | 335 |
|
30397 | 336 |
\item @{antiquotation_option_def unique_names}~@{text "= bool"} |
337 |
determines whether the printed version of qualified names should be |
|
338 |
made sufficiently long to avoid overlap with names declared further |
|
339 |
back. Set to @{text false} for more concise output. |
|
27043 | 340 |
|
30397 | 341 |
\item @{antiquotation_option_def eta_contract}~@{text "= bool"} |
342 |
prints terms in @{text \<eta>}-contracted form. |
|
27043 | 343 |
|
30397 | 344 |
\item @{antiquotation_option_def display}~@{text "= bool"} indicates |
345 |
if the text is to be output as multi-line ``display material'', |
|
346 |
rather than a small piece of text without line breaks (which is the |
|
347 |
default). |
|
27043 | 348 |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
349 |
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
|
350 |
the main theory text. |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
351 |
|
30397 | 352 |
\item @{antiquotation_option_def break}~@{text "= bool"} controls |
353 |
line breaks in non-display material. |
|
27043 | 354 |
|
30397 | 355 |
\item @{antiquotation_option_def quotes}~@{text "= bool"} indicates |
356 |
if the output should be enclosed in double quotes. |
|
27043 | 357 |
|
30397 | 358 |
\item @{antiquotation_option_def mode}~@{text "= name"} adds @{text |
359 |
name} to the print mode to be used for presentation. Note that the |
|
360 |
standard setup for {\LaTeX} output is already present by default, |
|
361 |
including the modes @{text latex} and @{text xsymbols}. |
|
27043 | 362 |
|
30397 | 363 |
\item @{antiquotation_option_def margin}~@{text "= nat"} and |
364 |
@{antiquotation_option_def indent}~@{text "= nat"} change the margin |
|
365 |
or indentation for pretty printing of display material. |
|
27043 | 366 |
|
30397 | 367 |
\item @{antiquotation_option_def goals_limit}~@{text "= nat"} |
368 |
determines the maximum number of goals to be printed (for goal-based |
|
369 |
antiquotation). |
|
27043 | 370 |
|
30397 | 371 |
\item @{antiquotation_option_def source}~@{text "= bool"} prints the |
372 |
original source text of the antiquotation arguments, rather than its |
|
373 |
internal representation. Note that formal checking of |
|
374 |
@{antiquotation "thm"}, @{antiquotation "term"}, etc. is still |
|
375 |
enabled; use the @{antiquotation "text"} antiquotation for unchecked |
|
376 |
output. |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
377 |
|
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
378 |
Regular @{text "term"} and @{text "typ"} antiquotations with @{text |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
379 |
"source = false"} involve a full round-trip from the original source |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
380 |
to an internalized logical entity back to a source form, according |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
381 |
to the syntax of the current context. Thus the printed output is |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
382 |
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
|
383 |
as the underlying theory is changed later on. |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
384 |
|
30397 | 385 |
In contrast, @{antiquotation_option_def source}~@{text "= true"} |
386 |
admits direct printing of the given source text, with the desirable |
|
387 |
well-formedness check in the background, but without modification of |
|
388 |
the printed text. |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
389 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
390 |
\end{description} |
27043 | 391 |
|
392 |
For boolean flags, ``@{text "name = true"}'' may be abbreviated as |
|
393 |
``@{text name}''. All of the above flags are disabled by default, |
|
28749
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
394 |
unless changed from ML, say in the @{verbatim "ROOT.ML"} of the |
99f6da3bbbf7
renamed "formal comments" to "document comments";
wenzelm
parents:
28747
diff
changeset
|
395 |
logic session. |
27043 | 396 |
*} |
397 |
||
398 |
||
28750 | 399 |
section {* Markup via command tags \label{sec:tags} *} |
27043 | 400 |
|
28750 | 401 |
text {* Each Isabelle/Isar command may be decorated by additional |
402 |
presentation tags, to indicate some modification in the way it is |
|
403 |
printed in the document. |
|
27043 | 404 |
|
405 |
\indexouternonterm{tags} |
|
406 |
\begin{rail} |
|
407 |
tags: ( tag * ) |
|
408 |
; |
|
409 |
tag: '\%' (ident | string) |
|
410 |
\end{rail} |
|
411 |
||
28750 | 412 |
Some tags are pre-declared for certain classes of commands, serving |
413 |
as default markup if no tags are given in the text: |
|
27043 | 414 |
|
28750 | 415 |
\medskip |
27043 | 416 |
\begin{tabular}{ll} |
417 |
@{text "theory"} & theory begin/end \\ |
|
418 |
@{text "proof"} & all proof commands \\ |
|
419 |
@{text "ML"} & all commands involving ML code \\ |
|
420 |
\end{tabular} |
|
421 |
||
28750 | 422 |
\medskip The Isabelle document preparation system |
423 |
\cite{isabelle-sys} allows tagged command regions to be presented |
|
27043 | 424 |
specifically, e.g.\ to fold proof texts, or drop parts of the text |
425 |
completely. |
|
426 |
||
28750 | 427 |
For example ``@{command "by"}~@{text "%invisible auto"}'' causes |
428 |
that piece of proof to be treated as @{text invisible} instead of |
|
429 |
@{text "proof"} (the default), which may be shown or hidden |
|
27043 | 430 |
depending on the document setup. In contrast, ``@{command |
28750 | 431 |
"by"}~@{text "%visible auto"}'' forces this text to be shown |
27043 | 432 |
invariably. |
433 |
||
434 |
Explicit tag specifications within a proof apply to all subsequent |
|
435 |
commands of the same level of nesting. For example, ``@{command |
|
28750 | 436 |
"proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' forces the whole |
437 |
sub-proof to be typeset as @{text visible} (unless some of its parts |
|
438 |
are tagged differently). |
|
439 |
||
440 |
\medskip Command tags merely produce certain markup environments for |
|
441 |
type-setting. The meaning of these is determined by {\LaTeX} |
|
442 |
macros, as defined in @{"file" "~~/lib/texinputs/isabelle.sty"} or |
|
443 |
by the document author. The Isabelle document preparation tools |
|
444 |
also provide some high-level options to specify the meaning of |
|
445 |
arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding |
|
446 |
parts of the text. Logic sessions may also specify ``document |
|
447 |
versions'', where given tags are interpreted in some particular way. |
|
448 |
Again see \cite{isabelle-sys} for further details. |
|
27043 | 449 |
*} |
450 |
||
451 |
||
452 |
section {* Draft presentation *} |
|
453 |
||
454 |
text {* |
|
455 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
456 |
@{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
457 |
@{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\ |
27043 | 458 |
\end{matharray} |
459 |
||
460 |
\begin{rail} |
|
461 |
('display\_drafts' | 'print\_drafts') (name +) |
|
462 |
; |
|
463 |
\end{rail} |
|
464 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
465 |
\begin{description} |
27043 | 466 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
467 |
\item @{command "display_drafts"}~@{text paths} and @{command |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
468 |
"print_drafts"}~@{text paths} perform simple output of a given list |
27043 | 469 |
of raw source files. Only those symbols that do not require |
470 |
additional {\LaTeX} packages are displayed properly, everything else |
|
471 |
is left verbatim. |
|
472 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
28750
diff
changeset
|
473 |
\end{description} |
27043 | 474 |
*} |
475 |
||
476 |
end |