| author | wenzelm |
| Mon, 21 Dec 2015 17:20:57 +0100 | |
| changeset 61889 | 42d902e074e8 |
| parent 61645 | ae5e55d03e45 |
| child 63414 | beb987127d0f |
| 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 |
| 61645 | 5 |
no_translations |
6 |
("prop") "P \<and> Q \<Longrightarrow> R" <= ("prop") "P \<Longrightarrow> Q \<Longrightarrow> R"
|
|
| 15337 | 7 |
(*>*) |
| 49239 | 8 |
text{*
|
9 |
\section{Introduction}
|
|
| 15337 | 10 |
|
| 49239 | 11 |
This document is for those Isabelle users who have mastered |
| 15337 | 12 |
the art of mixing \LaTeX\ text and Isabelle theories and never want to |
13 |
typeset a theorem by hand anymore because they have experienced the |
|
| 56977 | 14 |
bliss of writing \verb!@!\verb!{thm[display,mode=latex_sum] setsum_Suc_diff [no_vars]}!
|
| 15337 | 15 |
and seeing Isabelle typeset it for them: |
| 56977 | 16 |
@{thm[display,mode=latex_sum] setsum_Suc_diff[no_vars]}
|
| 15342 | 17 |
No typos, no omissions, no sweat. |
18 |
If you have not experienced that joy, read Chapter 4, \emph{Presenting
|
|
| 58620 | 19 |
Theories}, @{cite LNCS2283} first.
|
| 15337 | 20 |
|
21 |
If you have mastered the art of Isabelle's \emph{antiquotations},
|
|
22 |
i.e.\ things like the above \verb!@!\verb!{thm...}!, beware: in your vanity
|
|
23 |
you may be tempted to think that all readers of the stunning ps or pdf |
|
24 |
documents you can now produce at the drop of a hat will be struck with |
|
25 |
awe at the beauty unfolding in front of their eyes. Until one day you |
|
26 |
come across that very critical of readers known as the ``common referee''. |
|
27 |
He has the nasty habit of refusing to understand unfamiliar notation |
|
28 |
like Isabelle's infamous @{text"\<lbrakk> \<rbrakk> \<Longrightarrow>"} no matter how many times you
|
|
29 |
explain it in your paper. Even worse, he thinks that using @{text"\<lbrakk>
|
|
30 |
\<rbrakk>"} for anything other than denotational semantics is a cardinal sin |
|
| 15342 | 31 |
that must be punished by instant rejection. |
| 15337 | 32 |
|
33 |
||
34 |
This document shows you how to make Isabelle and \LaTeX\ cooperate to |
|
35 |
produce ordinary looking mathematics that hides the fact that it was |
|
| 15471 | 36 |
typeset by a machine. You merely need to load the right files: |
37 |
\begin{itemize}
|
|
38 |
\item Import theory \texttt{LaTeXsugar} in the header of your own
|
|
39 |
theory. You may also want bits of \texttt{OptionalSugar}, which you can
|
|
40 |
copy selectively into your own theory or import as a whole. Both |
|
41 |
theories live in \texttt{HOL/Library} and are found automatically.
|
|
| 15378 | 42 |
|
| 15471 | 43 |
\item Should you need additional \LaTeX\ packages (the text will tell |
44 |
you so), you include them at the beginning of your \LaTeX\ document, |
|
| 16153 | 45 |
typically in \texttt{root.tex}. For a start, you should
|
46 |
\verb!\usepackage{amssymb}! --- otherwise typesetting
|
|
47 |
@{prop[source]"\<not>(\<exists>x. P x)"} will fail because the AMS symbol
|
|
48 |
@{text"\<nexists>"} is missing.
|
|
| 15471 | 49 |
\end{itemize}
|
| 49239 | 50 |
|
| 15342 | 51 |
|
| 49239 | 52 |
\section{HOL syntax}
|
| 15342 | 53 |
|
| 49239 | 54 |
\subsection{Logic}
|
| 15342 | 55 |
|
| 49239 | 56 |
The formula @{prop[source]"\<not>(\<exists>x. P x)"} is typeset as @{prop"~(EX x. P x)"}.
|
| 16153 | 57 |
|
58 |
The predefined constructs @{text"if"}, @{text"let"} and
|
|
| 15342 | 59 |
@{text"case"} are set in sans serif font to distinguish them from
|
60 |
other functions. This improves readability: |
|
61 |
\begin{itemize}
|
|
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
49628
diff
changeset
|
62 |
\item @{term"if b then e\<^sub>1 else e\<^sub>2"} instead of @{text"if b then e\<^sub>1 else e\<^sub>2"}.
|
|
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
49628
diff
changeset
|
63 |
\item @{term"let x = e\<^sub>1 in e\<^sub>2"} instead of @{text"let x = e\<^sub>1 in e\<^sub>2"}.
|
|
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
49628
diff
changeset
|
64 |
\item @{term"case x of True \<Rightarrow> e\<^sub>1 | False \<Rightarrow> e\<^sub>2"} instead of\\
|
|
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
49628
diff
changeset
|
65 |
@{text"case x of True \<Rightarrow> e\<^sub>1 | False \<Rightarrow> e\<^sub>2"}.
|
| 15342 | 66 |
\end{itemize}
|
67 |
||
| 49239 | 68 |
\subsection{Sets}
|
| 15337 | 69 |
|
| 49239 | 70 |
Although set syntax in HOL is already close to |
| 15342 | 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))"}.
|
| 56977 | 77 |
\item @{term"card A"} instead of @{text"card A"}.
|
| 15342 | 78 |
\end{itemize}
|
| 49239 | 79 |
|
| 15342 | 80 |
|
| 49239 | 81 |
\subsection{Lists}
|
| 15342 | 82 |
|
| 49239 | 83 |
If lists are used heavily, the following notations increase readability: |
| 15342 | 84 |
\begin{itemize}
|
| 27688 | 85 |
\item @{term"x # xs"} instead of @{text"x # xs"},
|
86 |
where @{term"x # xs"} is also input syntax.
|
|
| 15342 | 87 |
\item @{term"length xs"} instead of @{text"length xs"}.
|
| 15385 | 88 |
\item @{term"nth xs n"} instead of @{text"nth xs n"},
|
| 15342 | 89 |
the $n$th element of @{text xs}.
|
90 |
||
| 22834 | 91 |
\item Human readers are good at converting automatically from lists to |
| 30502 | 92 |
sets. Hence \texttt{OptionalSugar} contains syntax for suppressing the
|
| 22834 | 93 |
conversion function @{const set}: for example, @{prop[source]"x \<in> set xs"}
|
94 |
becomes @{prop"x \<in> set xs"}.
|
|
95 |
||
| 15366 | 96 |
\item The @{text"@"} operation associates implicitly to the right,
|
97 |
which leads to unpleasant line breaks if the term is too long for one |
|
98 |
line. To avoid this, \texttt{OptionalSugar} contains syntax to group
|
|
99 |
@{text"@"}-terms to the left before printing, which leads to better
|
|
100 |
line breaking behaviour: |
|
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
49628
diff
changeset
|
101 |
@{term[display]"term\<^sub>0 @ term\<^sub>1 @ term\<^sub>2 @ term\<^sub>3 @ term\<^sub>4 @ term\<^sub>5 @ term\<^sub>6 @ term\<^sub>7 @ term\<^sub>8 @ term\<^sub>9 @ term\<^sub>1\<^sub>0"}
|
| 15366 | 102 |
|
| 15342 | 103 |
\end{itemize}
|
| 49239 | 104 |
|
| 15337 | 105 |
|
| 49239 | 106 |
\subsection{Numbers}
|
| 30502 | 107 |
|
| 49239 | 108 |
Coercions between numeric types are alien to mathematicians who |
| 30502 | 109 |
consider, for example, @{typ nat} as a subset of @{typ int}.
|
110 |
\texttt{OptionalSugar} contains syntax for suppressing numeric coercions such
|
|
111 |
as @{const int} @{text"::"} @{typ"nat \<Rightarrow> int"}. For example,
|
|
112 |
@{term[source]"int 5"} is printed as @{term "int 5"}. Embeddings of types
|
|
113 |
@{typ nat}, @{typ int}, @{typ real} are covered; non-injective coercions such
|
|
114 |
as @{const nat} @{text"::"} @{typ"int \<Rightarrow> nat"} are not and should not be
|
|
| 49239 | 115 |
hidden. |
| 30502 | 116 |
|
| 15337 | 117 |
|
| 49239 | 118 |
\section{Printing theorems}
|
| 15689 | 119 |
|
| 61645 | 120 |
The @{prop "P \<Longrightarrow> Q \<Longrightarrow> R"} syntax is a bit idiosyncratic. If you would like
|
121 |
to avoid it, you can easily print the premises as a conjunction: |
|
122 |
@{prop "P \<and> Q \<Longrightarrow> R"}. See \texttt{OptionalSugar} for the required ``code''.
|
|
123 |
||
| 49239 | 124 |
\subsection{Question marks}
|
125 |
||
126 |
If you print anything, especially theorems, containing |
|
| 15689 | 127 |
schematic variables they are prefixed with a question mark: |
128 |
\verb!@!\verb!{thm conjI}! results in @{thm conjI}. Most of the time
|
|
129 |
you would rather not see the question marks. There is an attribute |
|
130 |
\verb!no_vars! that you can attach to the theorem that turns its |
|
131 |
schematic into ordinary free variables: \verb!@!\verb!{thm conjI[no_vars]}!
|
|
132 |
results in @{thm conjI[no_vars]}.
|
|
133 |
||
134 |
This \verb!no_vars! business can become a bit tedious. |
|
135 |
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
|
136 |
\begin{quote}
|
| 49239 | 137 |
\verb!options [show_question_marks = false]! |
|
34877
ded5b770ec1c
formal antiquotations for ML snippets; no "open" unsynchronized references
haftmann
parents:
33323
diff
changeset
|
138 |
\end{quote}
|
| 49239 | 139 |
into the relevant \texttt{ROOT} file, just before the \texttt{theories} for that session.
|
| 33323 | 140 |
The rest of this document is produced with this flag set to \texttt{false}.
|
| 16075 | 141 |
|
| 49239 | 142 |
Hint: Setting \verb!show_question_marks! to \texttt{false} only
|
| 33323 | 143 |
suppresses question marks; variables that end in digits, |
144 |
e.g. @{text"x1"}, are still printed with a trailing @{text".0"},
|
|
145 |
e.g. @{text"x1.0"}, their internal index. This can be avoided by
|
|
| 61594 | 146 |
turning the last digit into a subscript: write \<^verbatim>\<open>x\<^sub>1\<close> and |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
49628
diff
changeset
|
147 |
obtain the much nicer @{text"x\<^sub>1"}. *}
|
| 15689 | 148 |
|
|
38980
af73cf0dc31f
turned show_question_marks into proper configuration option;
wenzelm
parents:
38798
diff
changeset
|
149 |
(*<*)declare [[show_question_marks = false]](*>*) |
| 15689 | 150 |
|
| 24496 | 151 |
subsection {*Qualified names*}
|
152 |
||
153 |
text{* If there are multiple declarations of the same name, Isabelle prints
|
|
154 |
the qualified name, for example @{text "T.length"}, where @{text T} is the
|
|
155 |
theory it is defined in, to distinguish it from the predefined @{const[source]
|
|
156 |
"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
|
157 |
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
|
158 |
configuration option in the context. |
| 49239 | 159 |
|
| 24496 | 160 |
|
| 49239 | 161 |
\subsection {Variable names\label{sec:varnames}}
|
| 16395 | 162 |
|
| 49239 | 163 |
It sometimes happens that you want to change the name of a |
| 16395 | 164 |
variable in a theorem before printing it. This can easily be achieved |
165 |
with the help of Isabelle's instantiation attribute \texttt{where}:
|
|
166 |
@{thm conjI[where P = \<phi> and Q = \<psi>]} is the result of
|
|
167 |
\begin{quote}
|
|
168 |
\verb!@!\verb!{thm conjI[where P = \<phi> and Q = \<psi>]}!
|
|
169 |
\end{quote}
|
|
170 |
To support the ``\_''-notation for irrelevant variables |
|
171 |
the constant \texttt{DUMMY} has been introduced:
|
|
|
55417
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
53015
diff
changeset
|
172 |
@{thm fst_conv[of _ DUMMY]} is produced by
|
| 16395 | 173 |
\begin{quote}
|
|
55417
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
53015
diff
changeset
|
174 |
\verb!@!\verb!{thm fst_conv[of _ DUMMY]}!
|
| 16395 | 175 |
\end{quote}
|
|
36138
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
176 |
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
|
177 |
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
|
178 |
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
|
179 |
\texttt{of} attribute:
|
|
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
180 |
\begin{quote}
|
|
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
181 |
\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
|
182 |
\end{quote}
|
|
1faa0fc34174
advertise [rename_abs] attribute in LaTeXsugar -- wish I had known about this earier.
krauss
parents:
34890
diff
changeset
|
183 |
produces @{thm split_paired_All[rename_abs _ l r]}.
|
| 49239 | 184 |
|
| 16395 | 185 |
|
| 49239 | 186 |
\subsection{Inference rules}
|
| 15337 | 187 |
|
| 49239 | 188 |
To print theorems as inference rules you need to include Didier |
| 58620 | 189 |
R\'emy's \texttt{mathpartir} package~@{cite mathpartir}
|
| 15342 | 190 |
for typesetting inference rules in your \LaTeX\ file. |
| 15337 | 191 |
|
| 15689 | 192 |
Writing \verb!@!\verb!{thm[mode=Rule] conjI}! produces
|
193 |
@{thm[mode=Rule] conjI}, even in the middle of a sentence.
|
|
| 15342 | 194 |
If you prefer your inference rule on a separate line, maybe with a name, |
195 |
\begin{center}
|
|
| 15689 | 196 |
@{thm[mode=Rule] conjI} {\sc conjI}
|
| 15342 | 197 |
\end{center}
|
198 |
is produced by |
|
| 15337 | 199 |
\begin{quote}
|
200 |
\verb!\begin{center}!\\
|
|
| 15689 | 201 |
\verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI}!\\
|
| 15337 | 202 |
\verb!\end{center}!
|
203 |
\end{quote}
|
|
| 24497 | 204 |
It is not recommended to use the standard \texttt{display} option
|
| 15342 | 205 |
together with \texttt{Rule} because centering does not work and because
|
206 |
the line breaking mechanisms of \texttt{display} and \texttt{mathpartir} can
|
|
207 |
clash. |
|
208 |
||
| 15337 | 209 |
Of course you can display multiple rules in this fashion: |
210 |
\begin{quote}
|
|
| 24497 | 211 |
\verb!\begin{center}!\\
|
| 15689 | 212 |
\verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI} \\[1ex]!\\
|
213 |
\verb!@!\verb!{thm[mode=Rule] conjE} {\sc disjI$_1$} \qquad!\\
|
|
214 |
\verb!@!\verb!{thm[mode=Rule] disjE} {\sc disjI$_2$}!\\
|
|
| 15337 | 215 |
\verb!\end{center}!
|
216 |
\end{quote}
|
|
217 |
yields |
|
| 24497 | 218 |
\begin{center}\small
|
| 15689 | 219 |
@{thm[mode=Rule] conjI} {\sc conjI} \\[1ex]
|
220 |
@{thm[mode=Rule] disjI1} {\sc disjI$_1$} \qquad
|
|
221 |
@{thm[mode=Rule] disjI2} {\sc disjI$_2$}
|
|
| 15337 | 222 |
\end{center}
|
223 |
||
| 15342 | 224 |
The \texttt{mathpartir} package copes well if there are too many
|
225 |
premises for one line: |
|
226 |
\begin{center}
|
|
227 |
@{prop[mode=Rule] "\<lbrakk> A \<longrightarrow> B; B \<longrightarrow> C; C \<longrightarrow> D; D \<longrightarrow> E; E \<longrightarrow> F; F \<longrightarrow> G;
|
|
228 |
G \<longrightarrow> H; H \<longrightarrow> I; I \<longrightarrow> J; J \<longrightarrow> K \<rbrakk> \<Longrightarrow> A \<longrightarrow> K"} |
|
229 |
\end{center}
|
|
230 |
||
| 15471 | 231 |
Limitations: 1. Premises and conclusion must each not be longer than |
232 |
the line. 2. Premises that are @{text"\<Longrightarrow>"}-implications are again
|
|
233 |
displayed with a horizontal line, which looks at least unusual. |
|
234 |
||
| 22329 | 235 |
|
236 |
In case you print theorems without premises no rule will be printed by the |
|
237 |
\texttt{Rule} print mode. However, you can use \texttt{Axiom} instead:
|
|
238 |
\begin{quote}
|
|
| 24497 | 239 |
\verb!\begin{center}!\\
|
| 22329 | 240 |
\verb!@!\verb!{thm[mode=Axiom] refl} {\sc refl}! \\
|
241 |
\verb!\end{center}!
|
|
242 |
\end{quote}
|
|
243 |
yields |
|
| 24497 | 244 |
\begin{center}
|
| 22329 | 245 |
@{thm[mode=Axiom] refl} {\sc refl}
|
246 |
\end{center}
|
|
| 49239 | 247 |
|
| 15342 | 248 |
|
| 49239 | 249 |
\subsection{Displays and font sizes}
|
| 24497 | 250 |
|
| 49239 | 251 |
When displaying theorems with the \texttt{display} option, for example as in
|
| 24497 | 252 |
\verb!@!\verb!{thm[display] refl}! @{thm[display] refl} the theorem is
|
253 |
set in small font. It uses the \LaTeX-macro \verb!\isastyle!, |
|
254 |
which is also the style that regular theory text is set in, e.g. *} |
|
255 |
||
256 |
lemma "t = t" |
|
257 |
(*<*)oops(*>*) |
|
258 |
||
259 |
text{* \noindent Otherwise \verb!\isastyleminor! is used,
|
|
260 |
which does not modify the font size (assuming you stick to the default |
|
261 |
\verb!\isabellestyle{it}! in \texttt{root.tex}). If you prefer
|
|
262 |
normal font size throughout your text, include |
|
263 |
\begin{quote}
|
|
264 |
\verb!\renewcommand{\isastyle}{\isastyleminor}!
|
|
265 |
\end{quote}
|
|
266 |
in \texttt{root.tex}. On the other hand, if you like the small font,
|
|
267 |
just put \verb!\isastyle! in front of the text in question, |
|
268 |
e.g.\ at the start of one of the center-environments above. |
|
269 |
||
270 |
The advantage of the display option is that you can display a whole |
|
271 |
list of theorems in one go. For example, |
|
|
46187
f009e0fe8643
updated example -- List.foldl is no longer defined via primrec;
wenzelm
parents:
42669
diff
changeset
|
272 |
\verb!@!\verb!{thm[display] append.simps}!
|
|
f009e0fe8643
updated example -- List.foldl is no longer defined via primrec;
wenzelm
parents:
42669
diff
changeset
|
273 |
generates @{thm[display] append.simps}
|
| 49239 | 274 |
|
| 24497 | 275 |
|
| 49239 | 276 |
\subsection{If-then}
|
| 15342 | 277 |
|
| 49239 | 278 |
If you prefer a fake ``natural language'' style you can produce |
| 15342 | 279 |
the body of |
280 |
\newtheorem{theorem}{Theorem}
|
|
281 |
\begin{theorem}
|
|
| 15689 | 282 |
@{thm[mode=IfThen] le_trans}
|
| 15342 | 283 |
\end{theorem}
|
284 |
by typing |
|
285 |
\begin{quote}
|
|
| 15689 | 286 |
\verb!@!\verb!{thm[mode=IfThen] le_trans}!
|
| 15342 | 287 |
\end{quote}
|
288 |
||
289 |
In order to prevent odd line breaks, the premises are put into boxes. |
|
290 |
At times this is too drastic: |
|
291 |
\begin{theorem}
|
|
292 |
@{prop[mode=IfThen] "longpremise \<Longrightarrow> longerpremise \<Longrightarrow> P(f(f(f(f(f(f(f(f(f(x)))))))))) \<Longrightarrow> longestpremise \<Longrightarrow> conclusion"}
|
|
293 |
\end{theorem}
|
|
| 16153 | 294 |
In which case you should use \texttt{IfThenNoBox} instead of
|
295 |
\texttt{IfThen}:
|
|
| 15342 | 296 |
\begin{theorem}
|
297 |
@{prop[mode=IfThenNoBox] "longpremise \<Longrightarrow> longerpremise \<Longrightarrow> P(f(f(f(f(f(f(f(f(f(x)))))))))) \<Longrightarrow> longestpremise \<Longrightarrow> conclusion"}
|
|
298 |
\end{theorem}
|
|
| 49239 | 299 |
|
| 15342 | 300 |
|
| 49239 | 301 |
\subsection{Doing it yourself\label{sec:yourself}}
|
| 16153 | 302 |
|
| 49239 | 303 |
If for some reason you want or need to present theorems your |
| 16153 | 304 |
own way, you can extract the premises and the conclusion explicitly |
305 |
and combine them as you like: |
|
306 |
\begin{itemize}
|
|
| 32891 | 307 |
\item \verb!@!\verb!{thm (prem 1)! $thm$\verb!}!
|
308 |
prints premise 1 of $thm$. |
|
309 |
\item \verb!@!\verb!{thm (concl)! $thm$\verb!}!
|
|
| 16153 | 310 |
prints the conclusion of $thm$. |
311 |
\end{itemize}
|
|
| 32891 | 312 |
For example, ``from @{thm (prem 2) conjI} and
|
313 |
@{thm (prem 1) conjI} we conclude @{thm (concl) conjI}''
|
|
| 16153 | 314 |
is produced by |
315 |
\begin{quote}
|
|
| 32891 | 316 |
\verb!from !\verb!@!\verb!{thm (prem 2) conjI}! \verb!and !\verb!@!\verb!{thm (prem 1) conjI}!\\
|
317 |
\verb!we conclude !\verb!@!\verb!{thm (concl) conjI}!
|
|
| 16153 | 318 |
\end{quote}
|
319 |
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
|
320 |
Styles like \verb!(prem 1)! are a general mechanism explained |
| 16153 | 321 |
in \S\ref{sec:styles}.
|
| 49239 | 322 |
|
323 |
||
324 |
\subsection{Patterns}
|
|
325 |
||
326 |
||
327 |
In \S\ref{sec:varnames} we shows how to create patterns containing ``@{term DUMMY}''.
|
|
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},
|
|
330 |
etc.\ are printed as patterns. \texttt{OptionalSugar} provides the following:
|
|
| 16153 | 331 |
|
| 49239 | 332 |
\begin{center}
|
333 |
\begin{tabular}{l@ {~~produced by~~}l}
|
|
334 |
@{term "let x = fst p in t"} & \verb!@!\verb!{term "let x = fst p in t"}!\\
|
|
335 |
@{term "let x = snd p in t"} & \verb!@!\verb!{term "let x = snd p in t"}!\\
|
|
336 |
@{term "let x = hd xs in t"} & \verb!@!\verb!{term "let x = hd xs in t"}!\\
|
|
337 |
@{term "let x = tl xs in t"} & \verb!@!\verb!{term "let x = tl xs in t"}!\\
|
|
338 |
@{term "let x = the y in t"} & \verb!@!\verb!{term "let x = the y in t"}!\\
|
|
339 |
\end{tabular}
|
|
340 |
\end{center}
|
|
341 |
||
342 |
||
343 |
\section {Styles\label{sec:styles}}
|
|
| 15366 | 344 |
|
| 49239 | 345 |
The \verb!thm! antiquotation works nicely for single theorems, but |
346 |
sets of equations as used in definitions are more difficult to |
|
347 |
typeset nicely: people tend to prefer aligned @{text "="} signs.
|
|
348 |
||
349 |
To deal with such cases where it is desirable to dive into the structure |
|
350 |
of terms and theorems, Isabelle offers antiquotations featuring ``styles'': |
|
351 |
||
352 |
\begin{quote}
|
|
353 |
\verb!@!\verb!{thm (style) thm}!\\
|
|
354 |
\verb!@!\verb!{prop (style) thm}!\\
|
|
355 |
\verb!@!\verb!{term (style) term}!\\
|
|
356 |
\verb!@!\verb!{term_type (style) term}!\\
|
|
357 |
\verb!@!\verb!{typeof (style) term}!\\
|
|
358 |
\end{quote}
|
|
| 15366 | 359 |
|
| 49239 | 360 |
A ``style'' is a transformation of a term. There are predefined |
361 |
styles, namely \verb!lhs! and \verb!rhs!, \verb!prem! with one argument, and \verb!concl!. |
|
362 |
For example, the output |
|
363 |
\begin{center}
|
|
364 |
\begin{tabular}{l@ {~~@{text "="}~~}l}
|
|
365 |
@{thm (lhs) append_Nil} & @{thm (rhs) append_Nil}\\
|
|
366 |
@{thm (lhs) append_Cons} & @{thm (rhs) append_Cons}
|
|
367 |
\end{tabular}
|
|
368 |
\end{center}
|
|
369 |
is produced by the following code: |
|
370 |
\begin{quote}
|
|
371 |
\verb!\begin{center}!\\
|
|
372 |
\verb!\begin{tabular}{l@ {~~!\verb!@!\verb!{text "="}~~}l}!\\
|
|
373 |
\verb!@!\verb!{thm (lhs) append_Nil} & @!\verb!{thm (rhs) append_Nil}\\!\\
|
|
374 |
\verb!@!\verb!{thm (lhs) append_Cons} & @!\verb!{thm (rhs) append_Cons}!\\
|
|
375 |
\verb!\end{tabular}!\\
|
|
376 |
\verb!\end{center}!
|
|
377 |
\end{quote}
|
|
378 |
Note the space between \verb!@! and \verb!{! in the tabular argument.
|
|
379 |
It prevents Isabelle from interpreting \verb!@ {~~...~~}!
|
|
380 |
as an antiquotation. The styles \verb!lhs! and \verb!rhs! |
|
381 |
extract the left hand side (or right hand side respectively) from the |
|
382 |
conclusion of propositions consisting of a binary operator |
|
383 |
(e.~g.~@{text "="}, @{text "\<equiv>"}, @{text "<"}).
|
|
| 15366 | 384 |
|
| 49239 | 385 |
Likewise, \verb!concl! may be used as a style to show just the |
386 |
conclusion of a proposition. For example, take \verb!hd_Cons_tl!: |
|
387 |
\begin{center}
|
|
388 |
@{thm hd_Cons_tl}
|
|
389 |
\end{center}
|
|
390 |
To print just the conclusion, |
|
391 |
\begin{center}
|
|
392 |
@{thm (concl) hd_Cons_tl}
|
|
393 |
\end{center}
|
|
394 |
type |
|
395 |
\begin{quote}
|
|
396 |
\verb!\begin{center}!\\
|
|
397 |
\verb!@!\verb!{thm (concl) hd_Cons_tl}!\\
|
|
398 |
\verb!\end{center}!
|
|
399 |
\end{quote}
|
|
400 |
Beware that any options must be placed \emph{before} the style, as in this example.
|
|
| 15366 | 401 |
|
| 49239 | 402 |
Further use cases can be found in \S\ref{sec:yourself}.
|
403 |
If you are not afraid of ML, you may also define your own styles. |
|
| 55837 | 404 |
Have a look at module @{ML_structure Term_Style}.
|
| 49239 | 405 |
|
406 |
||
407 |
\section {Proofs}
|
|
408 |
||
409 |
Full proofs, even if written in beautiful Isar style, are |
|
| 24497 | 410 |
likely to be too long and detailed to be included in conference |
411 |
papers, but some key lemmas might be of interest. |
|
412 |
It is usually easiest to put them in figures like the one in Fig.\ |
|
413 |
\ref{fig:proof}. This was achieved with the \isakeyword{text\_raw} command:
|
|
| 15366 | 414 |
*} |
415 |
text_raw {*
|
|
416 |
\begin{figure}
|
|
417 |
\begin{center}\begin{minipage}{0.6\textwidth}
|
|
| 24497 | 418 |
\isastyleminor\isamarkuptrue |
| 15366 | 419 |
*} |
420 |
lemma True |
|
421 |
proof - |
|
422 |
-- "pretty trivial" |
|
423 |
show True by force |
|
424 |
qed |
|
| 15428 | 425 |
text_raw {*
|
| 15366 | 426 |
\end{minipage}\end{center}
|
427 |
\caption{Example proof in a figure.}\label{fig:proof}
|
|
428 |
\end{figure}
|
|
429 |
*} |
|
430 |
text {*
|
|
431 |
||
432 |
\begin{quote}
|
|
433 |
\small |
|
434 |
\verb!text_raw {!\verb!*!\\
|
|
435 |
\verb! \begin{figure}!\\
|
|
436 |
\verb! \begin{center}\begin{minipage}{0.6\textwidth}!\\
|
|
| 24497 | 437 |
\verb! \isastyleminor\isamarkuptrue!\\ |
| 15366 | 438 |
\verb!*!\verb!}!\\ |
439 |
\verb!lemma True!\\ |
|
440 |
\verb!proof -!\\ |
|
441 |
\verb! -- "pretty trivial"!\\ |
|
442 |
\verb! show True by force!\\ |
|
443 |
\verb!qed!\\ |
|
444 |
\verb!text_raw {!\verb!*!\\
|
|
445 |
\verb! \end{minipage}\end{center}!\\
|
|
446 |
\verb! \caption{Example proof in a figure.}\label{fig:proof}!\\
|
|
447 |
\verb! \end{figure}!\\
|
|
448 |
\verb!*!\verb!}! |
|
449 |
\end{quote}
|
|
| 24497 | 450 |
|
451 |
Other theory text, e.g.\ definitions, can be put in figures, too. |
|
| 15342 | 452 |
|
| 49239 | 453 |
\section{Theory snippets}
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
454 |
|
| 49239 | 455 |
This section describes how to include snippets of a theory text in some other \LaTeX\ document. |
456 |
The typical scenario is that the description of your theory is not part of the theory text but |
|
457 |
a separate document that antiquotes bits of the theory. This works well for terms and theorems |
|
458 |
but there are no antiquotations, for example, for function definitions or proofs. Even if there are antiquotations, |
|
459 |
the output is usually a reformatted (by Isabelle) version of the input and may not look like |
|
460 |
you wanted it to look. Here is how to include a snippet of theory text (in \LaTeX\ form) in some |
|
461 |
other \LaTeX\ document, in 4 easy steps. Beware that these snippets are not processed by |
|
462 |
any antiquotation mechanism: the resulting \LaTeX\ text is more or less exactly what you wrote |
|
463 |
in the theory, without any added sugar. |
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
464 |
|
| 49239 | 465 |
\subsection{Theory markup}
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
466 |
|
| 49239 | 467 |
Include some markers at the beginning and the end of the theory snippet you want to cut out. |
468 |
You have to place the following lines before and after the snippet, where snippets must always be |
|
469 |
consecutive lines of theory text: |
|
470 |
\begin{quote}
|
|
471 |
\verb!\text_raw{!\verb!*\snip{!\emph{snippetname}\verb!}{1}{2}{%*!\verb!}!\\
|
|
472 |
\emph{theory text}\\
|
|
473 |
\verb!\text_raw{!\verb!*!\verb!}%endsnip*!\verb!}!
|
|
474 |
\end{quote}
|
|
475 |
where \emph{snippetname} should be a unique name for the snippet. The numbers \texttt{1}
|
|
476 |
and \texttt{2} are explained in a moment.
|
|
477 |
||
478 |
\subsection{Generate the \texttt{.tex} file}
|
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
479 |
|
| 49239 | 480 |
Run your theory \texttt{T} with the \texttt{isabelle} \texttt{build} tool
|
481 |
to generate the \LaTeX-file \texttt{T.tex} which is needed for the next step,
|
|
482 |
extraction of marked snippets. |
|
483 |
You may also want to process \texttt{T.tex} to generate a pdf document.
|
|
484 |
This requires a definition of \texttt{\char`\\snippet}:
|
|
485 |
\begin{verbatim}
|
|
486 |
\newcommand{\repeatisanl}[1]
|
|
487 |
{\ifnum#1=0\else\isanewline\repeatisanl{\numexpr#1-1}\fi}
|
|
488 |
\newcommand{\snip}[4]{\repeatisanl#2#4\repeatisanl#3}
|
|
489 |
\end{verbatim}
|
|
490 |
Parameter 2 and 3 of \texttt{\char`\\snippet} are numbers (the \texttt{1}
|
|
491 |
and \texttt{2} above) and determine how many newlines are inserted before and after the snippet.
|
|
492 |
Unfortunately \texttt{text\_raw} eats up all preceding and following newlines
|
|
493 |
and they have to be inserted again in this manner. Otherwise the document generated from \texttt{T.tex}
|
|
494 |
will look ugly around the snippets. It can take some iterations to get the number of required |
|
495 |
newlines exactly right. |
|
496 |
||
497 |
\subsection{Extract marked snippets}
|
|
498 |
\label{subsec:extract}
|
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
499 |
|
| 49239 | 500 |
Extract the marked bits of text with a shell-level script, e.g. |
501 |
\begin{quote}
|
|
502 |
\verb!sed -n '/\\snip{/,/endsnip/p' T.tex > !\emph{snippets}\verb!.tex!
|
|
503 |
\end{quote}
|
|
504 |
File \emph{snippets}\texttt{.tex} (the name is arbitrary) now contains a sequence of blocks like this
|
|
505 |
\begin{quote}
|
|
506 |
\verb!\snip{!\emph{snippetname}\verb!}{1}{2}{%!\\
|
|
507 |
\emph{theory text}\\
|
|
508 |
\verb!}%endsnip! |
|
509 |
\end{quote}
|
|
510 |
||
511 |
\subsection{Including snippets}
|
|
|
17031
ffa73448025e
added hint for position of aqu options in connection with styles
haftmann
parents:
16395
diff
changeset
|
512 |
|
| 49239 | 513 |
In the preamble of the document where the snippets are to be used you define \texttt{\char`\\snip}
|
514 |
and input \emph{snippets}\texttt{.tex}:
|
|
515 |
\begin{verbatim}
|
|
516 |
\newcommand{\snip}[4]
|
|
517 |
{\expandafter\newcommand\csname #1\endcsname{#4}}
|
|
518 |
\input{snippets}
|
|
519 |
\end{verbatim}
|
|
520 |
This definition of \texttt{\char`\\snip} simply has the effect of defining for each snippet
|
|
521 |
\emph{snippetname} a \LaTeX\ command \texttt{\char`\\}\emph{snippetname}
|
|
522 |
that produces the corresponding snippet text. In the body of your document you can display that text |
|
523 |
like this: |
|
524 |
\begin{quote}
|
|
525 |
\verb!\begin{isabelle}!\\
|
|
526 |
\texttt{\char`\\}\emph{snippetname}\\
|
|
527 |
\verb!\end{isabelle}!
|
|
528 |
\end{quote}
|
|
529 |
The \texttt{isabelle} environment is the one defined in the standard file
|
|
530 |
\texttt{isabelle.sty} which most likely you are loading anyway.
|
|
| 49628 | 531 |
|
532 |
||
533 |
\section{Antiquotation}
|
|
534 |
||
535 |
You want to show a constant and its type? Instead of going |
|
536 |
\verb!@!\verb!{const myconst}! \verb!@!\verb!{text "::"}! \verb!@!\verb!{typeof myconst}!,
|
|
537 |
you can just write \verb!@!\verb!{const_typ myconst}! using the new antiquotation
|
|
538 |
\texttt{const\_typ} defined in \texttt{LaTeXsugar}. For example,
|
|
539 |
\verb!@!\verb!{const_typ length}! produces @{const_typ length}.
|
|
540 |
||
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
541 |
*} |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
542 |
|
| 15337 | 543 |
(*<*) |
544 |
end |
|
| 16175 | 545 |
(*>*) |