author | blanchet |
Wed, 04 Jan 2012 00:30:53 +0100 (2012-01-03) | |
changeset 46108 | 1c314d838676 |
parent 42669 | 04dfffda5671 |
child 46187 | f009e0fe8643 |
permissions | -rw-r--r-- |
15337 | 1 |
(*<*) |
2 |
theory Sugar |
|
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
38980
diff
changeset
|
3 |
imports "~~/src/HOL/Library/LaTeXsugar" "~~/src/HOL/Library/OptionalSugar" |
15337 | 4 |
begin |
5 |
(*>*) |
|
6 |
||
7 |
section "Introduction" |
|
8 |
||
15385 | 9 |
text{* This document is for those Isabelle users who have mastered |
15337 | 10 |
the art of mixing \LaTeX\ text and Isabelle theories and never want to |
11 |
typeset a theorem by hand anymore because they have experienced the |
|
12 |
bliss of writing \verb!@!\verb!{thm[display]setsum_cartesian_product[no_vars]}! |
|
13 |
and seeing Isabelle typeset it for them: |
|
14 |
@{thm[display,eta_contract=false] setsum_cartesian_product[no_vars]} |
|
15342 | 15 |
No typos, no omissions, no sweat. |
16 |
If you have not experienced that joy, read Chapter 4, \emph{Presenting |
|
17 |
Theories}, \cite{LNCS2283} first. |
|
15337 | 18 |
|
19 |
If you have mastered the art of Isabelle's \emph{antiquotations}, |
|
20 |
i.e.\ things like the above \verb!@!\verb!{thm...}!, beware: in your vanity |
|
21 |
you may be tempted to think that all readers of the stunning ps or pdf |
|
22 |
documents you can now produce at the drop of a hat will be struck with |
|
23 |
awe at the beauty unfolding in front of their eyes. Until one day you |
|
24 |
come across that very critical of readers known as the ``common referee''. |
|
25 |
He has the nasty habit of refusing to understand unfamiliar notation |
|
26 |
like Isabelle's infamous @{text"\<lbrakk> \<rbrakk> \<Longrightarrow>"} no matter how many times you |
|
27 |
explain it in your paper. Even worse, he thinks that using @{text"\<lbrakk> |
|
28 |
\<rbrakk>"} for anything other than denotational semantics is a cardinal sin |
|
15342 | 29 |
that must be punished by instant rejection. |
15337 | 30 |
|
31 |
||
32 |
This document shows you how to make Isabelle and \LaTeX\ cooperate to |
|
33 |
produce ordinary looking mathematics that hides the fact that it was |
|
15471 | 34 |
typeset by a machine. You merely need to load the right files: |
35 |
\begin{itemize} |
|
36 |
\item Import theory \texttt{LaTeXsugar} in the header of your own |
|
37 |
theory. You may also want bits of \texttt{OptionalSugar}, which you can |
|
38 |
copy selectively into your own theory or import as a whole. Both |
|
39 |
theories live in \texttt{HOL/Library} and are found automatically. |
|
15378 | 40 |
|
15471 | 41 |
\item Should you need additional \LaTeX\ packages (the text will tell |
42 |
you so), you include them at the beginning of your \LaTeX\ document, |
|
16153 | 43 |
typically in \texttt{root.tex}. For a start, you should |
44 |
\verb!\usepackage{amssymb}! --- otherwise typesetting |
|
45 |
@{prop[source]"\<not>(\<exists>x. P x)"} will fail because the AMS symbol |
|
46 |
@{text"\<nexists>"} is missing. |
|
15471 | 47 |
\end{itemize} |
15342 | 48 |
*} |
49 |
||
50 |
section{* HOL syntax*} |
|
51 |
||
52 |
subsection{* Logic *} |
|
53 |
||
16153 | 54 |
text{* |
55 |
The formula @{prop[source]"\<not>(\<exists>x. P x)"} is typeset as @{prop"~(EX x. P x)"}. |
|
56 |
||
57 |
The predefined constructs @{text"if"}, @{text"let"} and |
|
15342 | 58 |
@{text"case"} are set in sans serif font to distinguish them from |
59 |
other functions. This improves readability: |
|
60 |
\begin{itemize} |
|
61 |
\item @{term"if b then e\<^isub>1 else e\<^isub>2"} instead of @{text"if b then e\<^isub>1 else e\<^isub>2"}. |
|
62 |
\item @{term"let x = e\<^isub>1 in e\<^isub>2"} instead of @{text"let x = e\<^isub>1 in e\<^isub>2"}. |
|
63 |
\item @{term"case x of True \<Rightarrow> e\<^isub>1 | False \<Rightarrow> e\<^isub>2"} instead of\\ |
|
64 |
@{text"case x of True \<Rightarrow> e\<^isub>1 | False \<Rightarrow> e\<^isub>2"}. |
|
65 |
\end{itemize} |
|
66 |
*} |
|
67 |
||
68 |
subsection{* Sets *} |
|
15337 | 69 |
|
15342 | 70 |
text{* Although set syntax in HOL is already close to |
71 |
standard, we provide a few further improvements: |
|
72 |
\begin{itemize} |
|
73 |
\item @{term"{x. P}"} instead of @{text"{x. P}"}. |
|
27688 | 74 |
\item @{term"{}"} instead of @{text"{}"}, where |
75 |
@{term"{}"} is also input syntax. |
|
15342 | 76 |
\item @{term"insert a (insert b (insert c M))"} instead of @{text"insert a (insert b (insert c M))"}. |
77 |
\end{itemize} |
|
78 |
*} |
|
79 |
||
80 |
subsection{* Lists *} |
|
81 |
||
82 |
text{* If lists are used heavily, the following notations increase readability: |
|
83 |
\begin{itemize} |
|
27688 | 84 |
\item @{term"x # xs"} instead of @{text"x # xs"}, |
85 |
where @{term"x # xs"} is also input syntax. |
|
15342 | 86 |
If you prefer more space around the $\cdot$ you have to redefine |
87 |
\verb!\isasymcdot! in \LaTeX: |
|
88 |
\verb!\renewcommand{\isasymcdot}{\isamath{\,\cdot\,}}! |
|
89 |
||
90 |
\item @{term"length xs"} instead of @{text"length xs"}. |
|
15385 | 91 |
\item @{term"nth xs n"} instead of @{text"nth xs n"}, |
15342 | 92 |
the $n$th element of @{text xs}. |
93 |
||
22834 | 94 |
\item Human readers are good at converting automatically from lists to |
30502 | 95 |
sets. Hence \texttt{OptionalSugar} contains syntax for suppressing the |
22834 | 96 |
conversion function @{const set}: for example, @{prop[source]"x \<in> set xs"} |
97 |
becomes @{prop"x \<in> set xs"}. |
|
98 |
||
15366 | 99 |
\item The @{text"@"} operation associates implicitly to the right, |
100 |
which leads to unpleasant line breaks if the term is too long for one |
|
101 |
line. To avoid this, \texttt{OptionalSugar} contains syntax to group |
|
102 |
@{text"@"}-terms to the left before printing, which leads to better |
|
103 |
line breaking behaviour: |
|
34878 | 104 |
@{term[display]"term\<^isub>0 @ term\<^isub>1 @ term\<^isub>2 @ term\<^isub>3 @ term\<^isub>4 @ term\<^isub>5 @ term\<^isub>6 @ term\<^isub>7 @ term\<^isub>8 @ term\<^isub>9 @ term\<^isub>1\<^isub>0"} |
15366 | 105 |
|
15342 | 106 |
\end{itemize} |
15337 | 107 |
*} |
108 |
||
30502 | 109 |
subsection{* Numbers *} |
110 |
||
111 |
text{* Coercions between numeric types are alien to mathematicians who |
|
112 |
consider, for example, @{typ nat} as a subset of @{typ int}. |
|
113 |
\texttt{OptionalSugar} contains syntax for suppressing numeric coercions such |
|
114 |
as @{const int} @{text"::"} @{typ"nat \<Rightarrow> int"}. For example, |
|
115 |
@{term[source]"int 5"} is printed as @{term "int 5"}. Embeddings of types |
|
116 |
@{typ nat}, @{typ int}, @{typ real} are covered; non-injective coercions such |
|
117 |
as @{const nat} @{text"::"} @{typ"int \<Rightarrow> nat"} are not and should not be |
|
118 |
hidden. *} |
|
119 |
||
15337 | 120 |
section "Printing theorems" |
121 |
||
15689 | 122 |
subsection "Question marks" |
123 |
||
124 |
text{* If you print anything, especially theorems, containing |
|
125 |
schematic variables they are prefixed with a question mark: |
|
126 |
\verb!@!\verb!{thm conjI}! results in @{thm conjI}. Most of the time |
|
127 |
you would rather not see the question marks. There is an attribute |
|
128 |
\verb!no_vars! that you can attach to the theorem that turns its |
|
129 |
schematic into ordinary free variables: \verb!@!\verb!{thm conjI[no_vars]}! |
|
130 |
results in @{thm conjI[no_vars]}. |
|
131 |
||
132 |
This \verb!no_vars! business can become a bit tedious. |
|
133 |
If you would rather never see question marks, simply put |
|
34877
ded5b770ec1c
formal antiquotations for ML snippets; no "open" unsynchronized references
haftmann
parents:
33323
diff
changeset
|
134 |
\begin{quote} |
42289
dafae095d733
discontinued special status of structure Printer;
wenzelm
parents:
41413
diff
changeset
|
135 |
@{ML "Printer.show_question_marks_default := false"}\verb!;! |
34877
ded5b770ec1c
formal antiquotations for ML snippets; no "open" unsynchronized references
haftmann
parents:
33323
diff
changeset
|
136 |
\end{quote} |
15689 | 137 |
at the beginning of your file \texttt{ROOT.ML}. |
33323 | 138 |
The rest of this document is produced with this flag set to \texttt{false}. |
16075 | 139 |
|
42289
dafae095d733
discontinued special status of structure Printer;
wenzelm
parents:
41413
diff
changeset
|
140 |
Hint: Setting @{ML Printer.show_question_marks_default} to \texttt{false} only |
33323 | 141 |
suppresses question marks; variables that end in digits, |
142 |
e.g. @{text"x1"}, are still printed with a trailing @{text".0"}, |
|
143 |
e.g. @{text"x1.0"}, their internal index. This can be avoided by |
|
144 |
turning the last digit into a subscript: write \verb!x\<^isub>1! and |
|
145 |
obtain the much nicer @{text"x\<^isub>1"}. *} |
|
15689 | 146 |
|
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
38798
diff
changeset
|
147 |
(*<*)declare [[show_question_marks = false]](*>*) |
15689 | 148 |
|
24496 | 149 |
subsection {*Qualified names*} |
150 |
||
151 |
text{* If there are multiple declarations of the same name, Isabelle prints |
|
152 |
the qualified name, for example @{text "T.length"}, where @{text T} is the |
|
153 |
theory it is defined in, to distinguish it from the predefined @{const[source] |
|
154 |
"List.length"}. In case there is no danger of confusion, you can insist on |
|
42669
04dfffda5671
more conventional naming scheme: names_long, names_short, names_unique;
wenzelm
parents:
42358
diff
changeset
|
155 |
short names (no qualifiers) by setting the \verb!names_short! |
42358
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents:
42289
diff
changeset
|
156 |
configuration option in the context. |
24496 | 157 |
*} |
158 |
||
17127 | 159 |
subsection {*Variable names\label{sec:varnames}*} |
16395 | 160 |
|
161 |
text{* It sometimes happens that you want to change the name of a |
|
162 |
variable in a theorem before printing it. This can easily be achieved |
|
163 |
with the help of Isabelle's instantiation attribute \texttt{where}: |
|
164 |
@{thm conjI[where P = \<phi> and Q = \<psi>]} is the result of |
|
165 |
\begin{quote} |
|
166 |
\verb!@!\verb!{thm conjI[where P = \<phi> and Q = \<psi>]}! |
|
167 |
\end{quote} |
|
168 |
To support the ``\_''-notation for irrelevant variables |
|
169 |
the constant \texttt{DUMMY} has been introduced: |
|
170 |
@{thm fst_conv[where b = DUMMY]} is produced by |
|
171 |
\begin{quote} |
|
172 |
\verb!@!\verb!{thm fst_conv[where b = DUMMY]}! |
|
173 |
\end{quote} |
|
36138
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
174 |
Variables that are bound by quantifiers or lambdas cannot be renamed |
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
175 |
like this. Instead, the attribute \texttt{rename\_abs} does the |
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
176 |
job. It expects a list of names or underscores, similar to the |
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
177 |
\texttt{of} attribute: |
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
178 |
\begin{quote} |
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
179 |
\verb!@!\verb!{thm split_paired_All[rename_abs _ l r]}! |
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
180 |
\end{quote} |
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
181 |
produces @{thm split_paired_All[rename_abs _ l r]}. |
16395 | 182 |
*} |
183 |
||
15337 | 184 |
subsection "Inference rules" |
185 |
||
15342 | 186 |
text{* To print theorems as inference rules you need to include Didier |
187 |
R\'emy's \texttt{mathpartir} package~\cite{mathpartir} |
|
188 |
for typesetting inference rules in your \LaTeX\ file. |
|
15337 | 189 |
|
15689 | 190 |
Writing \verb!@!\verb!{thm[mode=Rule] conjI}! produces |
191 |
@{thm[mode=Rule] conjI}, even in the middle of a sentence. |
|
15342 | 192 |
If you prefer your inference rule on a separate line, maybe with a name, |
193 |
\begin{center} |
|
15689 | 194 |
@{thm[mode=Rule] conjI} {\sc conjI} |
15342 | 195 |
\end{center} |
196 |
is produced by |
|
15337 | 197 |
\begin{quote} |
198 |
\verb!\begin{center}!\\ |
|
15689 | 199 |
\verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI}!\\ |
15337 | 200 |
\verb!\end{center}! |
201 |
\end{quote} |
|
24497 | 202 |
It is not recommended to use the standard \texttt{display} option |
15342 | 203 |
together with \texttt{Rule} because centering does not work and because |
204 |
the line breaking mechanisms of \texttt{display} and \texttt{mathpartir} can |
|
205 |
clash. |
|
206 |
||
15337 | 207 |
Of course you can display multiple rules in this fashion: |
208 |
\begin{quote} |
|
24497 | 209 |
\verb!\begin{center}!\\ |
15689 | 210 |
\verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI} \\[1ex]!\\ |
211 |
\verb!@!\verb!{thm[mode=Rule] conjE} {\sc disjI$_1$} \qquad!\\ |
|
212 |
\verb!@!\verb!{thm[mode=Rule] disjE} {\sc disjI$_2$}!\\ |
|
15337 | 213 |
\verb!\end{center}! |
214 |
\end{quote} |
|
215 |
yields |
|
24497 | 216 |
\begin{center}\small |
15689 | 217 |
@{thm[mode=Rule] conjI} {\sc conjI} \\[1ex] |
218 |
@{thm[mode=Rule] disjI1} {\sc disjI$_1$} \qquad |
|
219 |
@{thm[mode=Rule] disjI2} {\sc disjI$_2$} |
|
15337 | 220 |
\end{center} |
221 |
||
15342 | 222 |
The \texttt{mathpartir} package copes well if there are too many |
223 |
premises for one line: |
|
224 |
\begin{center} |
|
225 |
@{prop[mode=Rule] "\<lbrakk> A \<longrightarrow> B; B \<longrightarrow> C; C \<longrightarrow> D; D \<longrightarrow> E; E \<longrightarrow> F; F \<longrightarrow> G; |
|
226 |
G \<longrightarrow> H; H \<longrightarrow> I; I \<longrightarrow> J; J \<longrightarrow> K \<rbrakk> \<Longrightarrow> A \<longrightarrow> K"} |
|
227 |
\end{center} |
|
228 |
||
15471 | 229 |
Limitations: 1. Premises and conclusion must each not be longer than |
230 |
the line. 2. Premises that are @{text"\<Longrightarrow>"}-implications are again |
|
231 |
displayed with a horizontal line, which looks at least unusual. |
|
232 |
||
22329 | 233 |
|
234 |
In case you print theorems without premises no rule will be printed by the |
|
235 |
\texttt{Rule} print mode. However, you can use \texttt{Axiom} instead: |
|
236 |
\begin{quote} |
|
24497 | 237 |
\verb!\begin{center}!\\ |
22329 | 238 |
\verb!@!\verb!{thm[mode=Axiom] refl} {\sc refl}! \\ |
239 |
\verb!\end{center}! |
|
240 |
\end{quote} |
|
241 |
yields |
|
24497 | 242 |
\begin{center} |
22329 | 243 |
@{thm[mode=Axiom] refl} {\sc refl} |
244 |
\end{center} |
|
15337 | 245 |
*} |
15342 | 246 |
|
24497 | 247 |
subsection "Displays and font sizes" |
248 |
||
249 |
text{* When displaying theorems with the \texttt{display} option, e.g. |
|
250 |
\verb!@!\verb!{thm[display] refl}! @{thm[display] refl} the theorem is |
|
251 |
set in small font. It uses the \LaTeX-macro \verb!\isastyle!, |
|
252 |
which is also the style that regular theory text is set in, e.g. *} |
|
253 |
||
254 |
lemma "t = t" |
|
255 |
(*<*)oops(*>*) |
|
256 |
||
257 |
text{* \noindent Otherwise \verb!\isastyleminor! is used, |
|
258 |
which does not modify the font size (assuming you stick to the default |
|
259 |
\verb!\isabellestyle{it}! in \texttt{root.tex}). If you prefer |
|
260 |
normal font size throughout your text, include |
|
261 |
\begin{quote} |
|
262 |
\verb!\renewcommand{\isastyle}{\isastyleminor}! |
|
263 |
\end{quote} |
|
264 |
in \texttt{root.tex}. On the other hand, if you like the small font, |
|
265 |
just put \verb!\isastyle! in front of the text in question, |
|
266 |
e.g.\ at the start of one of the center-environments above. |
|
267 |
||
268 |
The advantage of the display option is that you can display a whole |
|
269 |
list of theorems in one go. For example, |
|
270 |
\verb!@!\verb!{thm[display] foldl.simps}! |
|
271 |
generates @{thm[display] foldl.simps} |
|
272 |
*} |
|
273 |
||
274 |
subsection "If-then" |
|
15342 | 275 |
|
276 |
text{* If you prefer a fake ``natural language'' style you can produce |
|
277 |
the body of |
|
278 |
\newtheorem{theorem}{Theorem} |
|
279 |
\begin{theorem} |
|
15689 | 280 |
@{thm[mode=IfThen] le_trans} |
15342 | 281 |
\end{theorem} |
282 |
by typing |
|
283 |
\begin{quote} |
|
15689 | 284 |
\verb!@!\verb!{thm[mode=IfThen] le_trans}! |
15342 | 285 |
\end{quote} |
286 |
||
287 |
In order to prevent odd line breaks, the premises are put into boxes. |
|
288 |
At times this is too drastic: |
|
289 |
\begin{theorem} |
|
290 |
@{prop[mode=IfThen] "longpremise \<Longrightarrow> longerpremise \<Longrightarrow> P(f(f(f(f(f(f(f(f(f(x)))))))))) \<Longrightarrow> longestpremise \<Longrightarrow> conclusion"} |
|
291 |
\end{theorem} |
|
16153 | 292 |
In which case you should use \texttt{IfThenNoBox} instead of |
293 |
\texttt{IfThen}: |
|
15342 | 294 |
\begin{theorem} |
295 |
@{prop[mode=IfThenNoBox] "longpremise \<Longrightarrow> longerpremise \<Longrightarrow> P(f(f(f(f(f(f(f(f(f(x)))))))))) \<Longrightarrow> longestpremise \<Longrightarrow> conclusion"} |
|
296 |
\end{theorem} |
|
15366 | 297 |
*} |
15342 | 298 |
|
16166 | 299 |
subsection{* Doing it yourself\label{sec:yourself}*} |
16153 | 300 |
|
301 |
text{* If for some reason you want or need to present theorems your |
|
302 |
own way, you can extract the premises and the conclusion explicitly |
|
303 |
and combine them as you like: |
|
304 |
\begin{itemize} |
|
32891 | 305 |
\item \verb!@!\verb!{thm (prem 1)! $thm$\verb!}! |
306 |
prints premise 1 of $thm$. |
|
307 |
\item \verb!@!\verb!{thm (concl)! $thm$\verb!}! |
|
16153 | 308 |
prints the conclusion of $thm$. |
309 |
\end{itemize} |
|
32891 | 310 |
For example, ``from @{thm (prem 2) conjI} and |
311 |
@{thm (prem 1) conjI} we conclude @{thm (concl) conjI}'' |
|
16153 | 312 |
is produced by |
313 |
\begin{quote} |
|
32891 | 314 |
\verb!from !\verb!@!\verb!{thm (prem 2) conjI}! \verb!and !\verb!@!\verb!{thm (prem 1) conjI}!\\ |
315 |
\verb!we conclude !\verb!@!\verb!{thm (concl) conjI}! |
|
16153 | 316 |
\end{quote} |
317 |
Thus you can rearrange or hide premises and typeset the theorem as you like. |
|
32898
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32891
diff
changeset
|
318 |
Styles like \verb!(prem 1)! are a general mechanism explained |
16153 | 319 |
in \S\ref{sec:styles}. |
320 |
*} |
|
321 |
||
15366 | 322 |
subsection "Patterns" |
323 |
||
324 |
text {* |
|
325 |
||
17127 | 326 |
In \S\ref{sec:varnames} we shows how to create patterns containing |
327 |
``@{term DUMMY}''. |
|
15366 | 328 |
You can drive this game even further and extend the syntax of let |
329 |
bindings such that certain functions like @{term fst}, @{term hd}, |
|
15368 | 330 |
etc.\ are printed as patterns. \texttt{OptionalSugar} provides the |
331 |
following: |
|
15366 | 332 |
|
333 |
\begin{center} |
|
334 |
\begin{tabular}{l@ {~~produced by~~}l} |
|
335 |
@{term "let x = fst p in t"} & \verb!@!\verb!{term "let x = fst p in t"}!\\ |
|
336 |
@{term "let x = snd p in t"} & \verb!@!\verb!{term "let x = snd p in t"}!\\ |
|
337 |
@{term "let x = hd xs in t"} & \verb!@!\verb!{term "let x = hd xs in t"}!\\ |
|
338 |
@{term "let x = tl xs in t"} & \verb!@!\verb!{term "let x = tl xs in t"}!\\ |
|
339 |
@{term "let x = the y in t"} & \verb!@!\verb!{term "let x = the y in t"}!\\ |
|
340 |
\end{tabular} |
|
341 |
\end{center} |
|
342 |
*} |
|
343 |
||
16155 | 344 |
section "Proofs" |
15366 | 345 |
|
24497 | 346 |
text {* Full proofs, even if written in beautiful Isar style, are |
347 |
likely to be too long and detailed to be included in conference |
|
348 |
papers, but some key lemmas might be of interest. |
|
349 |
It is usually easiest to put them in figures like the one in Fig.\ |
|
350 |
\ref{fig:proof}. This was achieved with the \isakeyword{text\_raw} command: |
|
15366 | 351 |
*} |
352 |
text_raw {* |
|
353 |
\begin{figure} |
|
354 |
\begin{center}\begin{minipage}{0.6\textwidth} |
|
24497 | 355 |
\isastyleminor\isamarkuptrue |
15366 | 356 |
*} |
357 |
lemma True |
|
358 |
proof - |
|
359 |
-- "pretty trivial" |
|
360 |
show True by force |
|
361 |
qed |
|
15428 | 362 |
text_raw {* |
15366 | 363 |
\end{minipage}\end{center} |
364 |
\caption{Example proof in a figure.}\label{fig:proof} |
|
365 |
\end{figure} |
|
366 |
*} |
|
367 |
text {* |
|
368 |
||
369 |
\begin{quote} |
|
370 |
\small |
|
371 |
\verb!text_raw {!\verb!*!\\ |
|
372 |
\verb! \begin{figure}!\\ |
|
373 |
\verb! \begin{center}\begin{minipage}{0.6\textwidth}!\\ |
|
24497 | 374 |
\verb! \isastyleminor\isamarkuptrue!\\ |
15366 | 375 |
\verb!*!\verb!}!\\ |
376 |
\verb!lemma True!\\ |
|
377 |
\verb!proof -!\\ |
|
378 |
\verb! -- "pretty trivial"!\\ |
|
379 |
\verb! show True by force!\\ |
|
380 |
\verb!qed!\\ |
|
381 |
\verb!text_raw {!\verb!*!\\ |
|
382 |
\verb! \end{minipage}\end{center}!\\ |
|
383 |
\verb! \caption{Example proof in a figure.}\label{fig:proof}!\\ |
|
384 |
\verb! \end{figure}!\\ |
|
385 |
\verb!*!\verb!}! |
|
386 |
\end{quote} |
|
24497 | 387 |
|
388 |
Other theory text, e.g.\ definitions, can be put in figures, too. |
|
15342 | 389 |
*} |
390 |
||
16155 | 391 |
section {*Styles\label{sec:styles}*} |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
392 |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
393 |
text {* |
15953 | 394 |
The \verb!thm! antiquotation works nicely for single theorems, but |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
395 |
sets of equations as used in definitions are more difficult to |
16040 | 396 |
typeset nicely: people tend to prefer aligned @{text "="} signs. |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
397 |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
398 |
To deal with such cases where it is desirable to dive into the structure |
16040 | 399 |
of terms and theorems, Isabelle offers antiquotations featuring |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
400 |
``styles'': |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
401 |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
402 |
\begin{quote} |
32891 | 403 |
\verb!@!\verb!{thm (style) thm}!\\ |
404 |
\verb!@!\verb!{prop (style) thm}!\\ |
|
32898
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32891
diff
changeset
|
405 |
\verb!@!\verb!{term (style) term}!\\ |
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32891
diff
changeset
|
406 |
\verb!@!\verb!{term_type (style) term}!\\ |
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32891
diff
changeset
|
407 |
\verb!@!\verb!{typeof (style) term}!\\ |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
408 |
\end{quote} |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
409 |
|
32898
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32891
diff
changeset
|
410 |
A ``style'' is a transformation of a term. There are predefined |
32891 | 411 |
styles, namely \verb!lhs! and \verb!rhs!, \verb!prem! with one argument, and \verb!concl!. |
16166 | 412 |
For example, |
16076 | 413 |
the output |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
414 |
\begin{center} |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
415 |
\begin{tabular}{l@ {~~@{text "="}~~}l} |
32891 | 416 |
@{thm (lhs) foldl_Nil} & @{thm (rhs) foldl_Nil}\\ |
417 |
@{thm (lhs) foldl_Cons} & @{thm (rhs) foldl_Cons} |
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
418 |
\end{tabular} |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
419 |
\end{center} |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
420 |
is produced by the following code: |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
421 |
\begin{quote} |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
422 |
\verb!\begin{center}!\\ |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
423 |
\verb!\begin{tabular}{l@ {~~!\verb!@!\verb!{text "="}~~}l}!\\ |
33323 | 424 |
\verb!@!\verb!{thm (lhs) foldl_Nil} & @!\verb!{thm (rhs) foldl_Nil}\\!\\ |
32891 | 425 |
\verb!@!\verb!{thm (lhs) foldl_Cons} & @!\verb!{thm (rhs) foldl_Cons}!\\ |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
426 |
\verb!\end{tabular}!\\ |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
427 |
\verb!\end{center}! |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
428 |
\end{quote} |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
429 |
Note the space between \verb!@! and \verb!{! in the tabular argument. |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
430 |
It prevents Isabelle from interpreting \verb!@ {~~...~~}! |
16076 | 431 |
as an antiquotation. The styles \verb!lhs! and \verb!rhs! |
27093 | 432 |
extract the left hand side (or right hand side respectively) from the |
16076 | 433 |
conclusion of propositions consisting of a binary operator |
16040 | 434 |
(e.~g.~@{text "="}, @{text "\<equiv>"}, @{text "<"}). |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
435 |
|
16165 | 436 |
Likewise, \verb!concl! may be used as a style to show just the |
16076 | 437 |
conclusion of a proposition. For example, take \verb!hd_Cons_tl!: |
16040 | 438 |
\begin{center} |
33323 | 439 |
@{thm hd_Cons_tl} |
16040 | 440 |
\end{center} |
441 |
To print just the conclusion, |
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
442 |
\begin{center} |
33323 | 443 |
@{thm (concl) hd_Cons_tl} |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
444 |
\end{center} |
16040 | 445 |
type |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
446 |
\begin{quote} |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
447 |
\verb!\begin{center}!\\ |
33323 | 448 |
\verb!@!\verb!{thm (concl) hd_Cons_tl}!\\ |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
449 |
\verb!\end{center}! |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
450 |
\end{quote} |
17127 | 451 |
Beware that any options must be placed \emph{before} |
32898
e871d897969c
term styles also cover antiquotations term_type and typeof
haftmann
parents:
32891
diff
changeset
|
452 |
the style, as in this example. |
17031
ffa73448025e
added hint for position of aqu options in connection with styles
haftmann
parents:
16395
diff
changeset
|
453 |
|
16166 | 454 |
Further use cases can be found in \S\ref{sec:yourself}. |
15953 | 455 |
If you are not afraid of ML, you may also define your own styles. |
32891 | 456 |
Have a look at module @{ML_struct Term_Style}. |
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
457 |
*} |
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
458 |
|
15337 | 459 |
(*<*) |
460 |
end |
|
16175 | 461 |
(*>*) |