| author | haftmann |
| Wed, 21 Mar 2007 08:07:19 +0100 | |
| changeset 22490 | 1822ec4fcecd |
| parent 22329 | e4325ce4e0c4 |
| child 22834 | bf67f798f063 |
| permissions | -rw-r--r-- |
| 15337 | 1 |
(*<*) |
2 |
theory Sugar |
|
| 15366 | 3 |
imports LaTeXsugar 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}"}.
|
|
74 |
\item @{term"{}"} instead of @{text"{}"}.
|
|
75 |
\item @{term"insert a (insert b (insert c M))"} instead of @{text"insert a (insert b (insert c M))"}.
|
|
76 |
\end{itemize}
|
|
77 |
*} |
|
78 |
||
79 |
subsection{* Lists *}
|
|
80 |
||
81 |
text{* If lists are used heavily, the following notations increase readability:
|
|
82 |
\begin{itemize}
|
|
83 |
\item @{term"x # xs"} instead of @{text"x # xs"}.
|
|
84 |
Exceptionally, @{term"x # xs"} is also input syntax.
|
|
85 |
If you prefer more space around the $\cdot$ you have to redefine |
|
86 |
\verb!\isasymcdot! in \LaTeX: |
|
87 |
\verb!\renewcommand{\isasymcdot}{\isamath{\,\cdot\,}}!
|
|
88 |
||
89 |
\item @{term"length xs"} instead of @{text"length xs"}.
|
|
| 15385 | 90 |
\item @{term"nth xs n"} instead of @{text"nth xs n"},
|
| 15342 | 91 |
the $n$th element of @{text xs}.
|
92 |
||
| 15366 | 93 |
\item The @{text"@"} operation associates implicitly to the right,
|
94 |
which leads to unpleasant line breaks if the term is too long for one |
|
95 |
line. To avoid this, \texttt{OptionalSugar} contains syntax to group
|
|
96 |
@{text"@"}-terms to the left before printing, which leads to better
|
|
97 |
line breaking behaviour: |
|
| 15673 | 98 |
@{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 | 99 |
|
| 15342 | 100 |
\end{itemize}
|
| 15337 | 101 |
*} |
102 |
||
103 |
section "Printing theorems" |
|
104 |
||
| 15689 | 105 |
subsection "Question marks" |
106 |
||
107 |
text{* If you print anything, especially theorems, containing
|
|
108 |
schematic variables they are prefixed with a question mark: |
|
109 |
\verb!@!\verb!{thm conjI}! results in @{thm conjI}. Most of the time
|
|
110 |
you would rather not see the question marks. There is an attribute |
|
111 |
\verb!no_vars! that you can attach to the theorem that turns its |
|
112 |
schematic into ordinary free variables: \verb!@!\verb!{thm conjI[no_vars]}!
|
|
113 |
results in @{thm conjI[no_vars]}.
|
|
114 |
||
115 |
This \verb!no_vars! business can become a bit tedious. |
|
116 |
If you would rather never see question marks, simply put |
|
117 |
\begin{verbatim}
|
|
| 15983 | 118 |
reset show_question_marks; |
| 15689 | 119 |
\end{verbatim}
|
120 |
at the beginning of your file \texttt{ROOT.ML}.
|
|
121 |
The rest of this document is produced with this flag reset. |
|
| 16075 | 122 |
|
123 |
Hint: Resetting \verb!show_question_marks! only supresses question |
|
124 |
marks; variables that end in digits, e.g. @{text"x1"}, are still
|
|
125 |
printed with a trailing @{text".0"}, e.g. @{text"x1.0"}, their
|
|
126 |
internal index. This can be avoided by turning the last digit into a |
|
127 |
subscript: write \verb!x\<^isub>1! and obtain the much nicer @{text"x\<^isub>1"}. *}
|
|
| 15689 | 128 |
|
| 15983 | 129 |
(*<*)ML"reset show_question_marks"(*>*) |
| 15689 | 130 |
|
| 17127 | 131 |
subsection {*Variable names\label{sec:varnames}*}
|
| 16395 | 132 |
|
133 |
text{* It sometimes happens that you want to change the name of a
|
|
134 |
variable in a theorem before printing it. This can easily be achieved |
|
135 |
with the help of Isabelle's instantiation attribute \texttt{where}:
|
|
136 |
@{thm conjI[where P = \<phi> and Q = \<psi>]} is the result of
|
|
137 |
\begin{quote}
|
|
138 |
\verb!@!\verb!{thm conjI[where P = \<phi> and Q = \<psi>]}!
|
|
139 |
\end{quote}
|
|
140 |
To support the ``\_''-notation for irrelevant variables |
|
141 |
the constant \texttt{DUMMY} has been introduced:
|
|
142 |
@{thm fst_conv[where b = DUMMY]} is produced by
|
|
143 |
\begin{quote}
|
|
144 |
\verb!@!\verb!{thm fst_conv[where b = DUMMY]}!
|
|
145 |
\end{quote}
|
|
146 |
*} |
|
147 |
||
| 15337 | 148 |
subsection "Inference rules" |
149 |
||
| 15342 | 150 |
text{* To print theorems as inference rules you need to include Didier
|
151 |
R\'emy's \texttt{mathpartir} package~\cite{mathpartir}
|
|
152 |
for typesetting inference rules in your \LaTeX\ file. |
|
| 15337 | 153 |
|
| 15689 | 154 |
Writing \verb!@!\verb!{thm[mode=Rule] conjI}! produces
|
155 |
@{thm[mode=Rule] conjI}, even in the middle of a sentence.
|
|
| 15342 | 156 |
If you prefer your inference rule on a separate line, maybe with a name, |
157 |
\begin{center}
|
|
| 15689 | 158 |
@{thm[mode=Rule] conjI} {\sc conjI}
|
| 15342 | 159 |
\end{center}
|
160 |
is produced by |
|
| 15337 | 161 |
\begin{quote}
|
162 |
\verb!\begin{center}!\\
|
|
| 15689 | 163 |
\verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI}!\\
|
| 15337 | 164 |
\verb!\end{center}!
|
165 |
\end{quote}
|
|
| 15342 | 166 |
It is not recommended to use the standard \texttt{display} attribute
|
167 |
together with \texttt{Rule} because centering does not work and because
|
|
168 |
the line breaking mechanisms of \texttt{display} and \texttt{mathpartir} can
|
|
169 |
clash. |
|
170 |
||
| 15337 | 171 |
Of course you can display multiple rules in this fashion: |
172 |
\begin{quote}
|
|
173 |
\verb!\begin{center}\isastyle!\\
|
|
| 15689 | 174 |
\verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI} \\[1ex]!\\
|
175 |
\verb!@!\verb!{thm[mode=Rule] conjE} {\sc disjI$_1$} \qquad!\\
|
|
176 |
\verb!@!\verb!{thm[mode=Rule] disjE} {\sc disjI$_2$}!\\
|
|
| 15337 | 177 |
\verb!\end{center}!
|
178 |
\end{quote}
|
|
179 |
yields |
|
180 |
\begin{center}\isastyle
|
|
| 15689 | 181 |
@{thm[mode=Rule] conjI} {\sc conjI} \\[1ex]
|
182 |
@{thm[mode=Rule] disjI1} {\sc disjI$_1$} \qquad
|
|
183 |
@{thm[mode=Rule] disjI2} {\sc disjI$_2$}
|
|
| 15337 | 184 |
\end{center}
|
185 |
Note that we included \verb!\isastyle! to obtain |
|
186 |
the smaller font that otherwise comes only with \texttt{display}.
|
|
187 |
||
| 15342 | 188 |
The \texttt{mathpartir} package copes well if there are too many
|
189 |
premises for one line: |
|
190 |
\begin{center}
|
|
191 |
@{prop[mode=Rule] "\<lbrakk> A \<longrightarrow> B; B \<longrightarrow> C; C \<longrightarrow> D; D \<longrightarrow> E; E \<longrightarrow> F; F \<longrightarrow> G;
|
|
192 |
G \<longrightarrow> H; H \<longrightarrow> I; I \<longrightarrow> J; J \<longrightarrow> K \<rbrakk> \<Longrightarrow> A \<longrightarrow> K"} |
|
193 |
\end{center}
|
|
194 |
||
| 15471 | 195 |
Limitations: 1. Premises and conclusion must each not be longer than |
196 |
the line. 2. Premises that are @{text"\<Longrightarrow>"}-implications are again
|
|
197 |
displayed with a horizontal line, which looks at least unusual. |
|
198 |
||
| 22329 | 199 |
|
200 |
In case you print theorems without premises no rule will be printed by the |
|
201 |
\texttt{Rule} print mode. However, you can use \texttt{Axiom} instead:
|
|
202 |
\begin{quote}
|
|
203 |
\verb!\begin{center}\isastyle!\\
|
|
204 |
\verb!@!\verb!{thm[mode=Axiom] refl} {\sc refl}! \\
|
|
205 |
\verb!\end{center}!
|
|
206 |
\end{quote}
|
|
207 |
yields |
|
208 |
\begin{center}\isastyle
|
|
209 |
@{thm[mode=Axiom] refl} {\sc refl}
|
|
210 |
\end{center}
|
|
211 |
||
212 |
||
213 |
||
| 15337 | 214 |
*} |
| 15342 | 215 |
|
216 |
subsection{*If-then*}
|
|
217 |
||
218 |
text{* If you prefer a fake ``natural language'' style you can produce
|
|
219 |
the body of |
|
220 |
\newtheorem{theorem}{Theorem}
|
|
221 |
\begin{theorem}
|
|
| 15689 | 222 |
@{thm[mode=IfThen] le_trans}
|
| 15342 | 223 |
\end{theorem}
|
224 |
by typing |
|
225 |
\begin{quote}
|
|
| 15689 | 226 |
\verb!@!\verb!{thm[mode=IfThen] le_trans}!
|
| 15342 | 227 |
\end{quote}
|
228 |
||
229 |
In order to prevent odd line breaks, the premises are put into boxes. |
|
230 |
At times this is too drastic: |
|
231 |
\begin{theorem}
|
|
232 |
@{prop[mode=IfThen] "longpremise \<Longrightarrow> longerpremise \<Longrightarrow> P(f(f(f(f(f(f(f(f(f(x)))))))))) \<Longrightarrow> longestpremise \<Longrightarrow> conclusion"}
|
|
233 |
\end{theorem}
|
|
| 16153 | 234 |
In which case you should use \texttt{IfThenNoBox} instead of
|
235 |
\texttt{IfThen}:
|
|
| 15342 | 236 |
\begin{theorem}
|
237 |
@{prop[mode=IfThenNoBox] "longpremise \<Longrightarrow> longerpremise \<Longrightarrow> P(f(f(f(f(f(f(f(f(f(x)))))))))) \<Longrightarrow> longestpremise \<Longrightarrow> conclusion"}
|
|
238 |
\end{theorem}
|
|
| 15366 | 239 |
*} |
| 15342 | 240 |
|
| 16166 | 241 |
subsection{* Doing it yourself\label{sec:yourself}*}
|
| 16153 | 242 |
|
243 |
text{* If for some reason you want or need to present theorems your
|
|
244 |
own way, you can extract the premises and the conclusion explicitly |
|
245 |
and combine them as you like: |
|
246 |
\begin{itemize}
|
|
| 16167 | 247 |
\item \verb!@!\verb!{thm_style prem1! $thm$\verb!}!
|
248 |
prints premise 1 of $thm$ (and similarly up to \texttt{prem9}).
|
|
| 16165 | 249 |
\item \verb!@!\verb!{thm_style concl! $thm$\verb!}!
|
| 16153 | 250 |
prints the conclusion of $thm$. |
251 |
\end{itemize}
|
|
| 16167 | 252 |
For example, ``from @{thm_style prem2 conjI} and
|
253 |
@{thm_style prem1 conjI} we conclude @{thm_style concl conjI}''
|
|
| 16153 | 254 |
is produced by |
255 |
\begin{quote}
|
|
| 16175 | 256 |
\verb!from !\verb!@!\verb!{thm_style prem2 conjI}! \verb!and !\verb!@!\verb!{thm_style prem1 conjI}!\\
|
| 16165 | 257 |
\verb!we conclude !\verb!@!\verb!{thm_style concl conjI}!
|
| 16153 | 258 |
\end{quote}
|
259 |
Thus you can rearrange or hide premises and typeset the theorem as you like. |
|
260 |
The \verb!thm_style! antiquotation is a general mechanism explained |
|
261 |
in \S\ref{sec:styles}.
|
|
262 |
*} |
|
263 |
||
| 15366 | 264 |
subsection "Patterns" |
265 |
||
266 |
text {*
|
|
267 |
||
| 17127 | 268 |
In \S\ref{sec:varnames} we shows how to create patterns containing
|
269 |
``@{term DUMMY}''.
|
|
| 15366 | 270 |
You can drive this game even further and extend the syntax of let |
271 |
bindings such that certain functions like @{term fst}, @{term hd},
|
|
| 15368 | 272 |
etc.\ are printed as patterns. \texttt{OptionalSugar} provides the
|
273 |
following: |
|
| 15366 | 274 |
|
275 |
\begin{center}
|
|
276 |
\begin{tabular}{l@ {~~produced by~~}l}
|
|
277 |
@{term "let x = fst p in t"} & \verb!@!\verb!{term "let x = fst p in t"}!\\
|
|
278 |
@{term "let x = snd p in t"} & \verb!@!\verb!{term "let x = snd p in t"}!\\
|
|
279 |
@{term "let x = hd xs in t"} & \verb!@!\verb!{term "let x = hd xs in t"}!\\
|
|
280 |
@{term "let x = tl xs in t"} & \verb!@!\verb!{term "let x = tl xs in t"}!\\
|
|
281 |
@{term "let x = the y in t"} & \verb!@!\verb!{term "let x = the y in t"}!\\
|
|
282 |
\end{tabular}
|
|
283 |
\end{center}
|
|
284 |
*} |
|
285 |
||
| 16155 | 286 |
section "Proofs" |
| 15366 | 287 |
|
288 |
text {*
|
|
| 15367 | 289 |
Full proofs, even if written in beautiful Isar style, are likely to |
| 15366 | 290 |
be too long and detailed to be included in conference papers, but |
291 |
some key lemmas might be of interest. |
|
292 |
||
293 |
It is usually easiest to put them in figures like the one in Fig.\ |
|
294 |
\ref{fig:proof}. This was achieved with the \isakeyword{text\_raw}
|
|
295 |
command: |
|
296 |
*} |
|
297 |
text_raw {*
|
|
298 |
\begin{figure}
|
|
299 |
\begin{center}\begin{minipage}{0.6\textwidth}
|
|
| 15428 | 300 |
\isastyle\isamarkuptrue |
| 15366 | 301 |
*} |
302 |
lemma True |
|
303 |
proof - |
|
304 |
-- "pretty trivial" |
|
305 |
show True by force |
|
306 |
qed |
|
| 15428 | 307 |
text_raw {*
|
| 15366 | 308 |
\end{minipage}\end{center}
|
309 |
\caption{Example proof in a figure.}\label{fig:proof}
|
|
310 |
\end{figure}
|
|
311 |
*} |
|
312 |
text {*
|
|
313 |
||
314 |
\begin{quote}
|
|
315 |
\small |
|
316 |
\verb!text_raw {!\verb!*!\\
|
|
317 |
\verb! \begin{figure}!\\
|
|
318 |
\verb! \begin{center}\begin{minipage}{0.6\textwidth}!\\
|
|
| 15428 | 319 |
\verb! \isastyle\isamarkuptrue!\\ |
| 15366 | 320 |
\verb!*!\verb!}!\\ |
321 |
\verb!lemma True!\\ |
|
322 |
\verb!proof -!\\ |
|
323 |
\verb! -- "pretty trivial"!\\ |
|
324 |
\verb! show True by force!\\ |
|
325 |
\verb!qed!\\ |
|
326 |
\verb!text_raw {!\verb!*!\\
|
|
327 |
\verb! \end{minipage}\end{center}!\\
|
|
328 |
\verb! \caption{Example proof in a figure.}\label{fig:proof}!\\
|
|
329 |
\verb! \end{figure}!\\
|
|
330 |
\verb!*!\verb!}! |
|
331 |
\end{quote}
|
|
332 |
||
| 15342 | 333 |
*} |
334 |
||
| 16155 | 335 |
section {*Styles\label{sec:styles}*}
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
336 |
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
337 |
text {*
|
| 15953 | 338 |
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
|
339 |
sets of equations as used in definitions are more difficult to |
| 16040 | 340 |
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
|
341 |
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
342 |
To deal with such cases where it is desirable to dive into the structure |
| 16040 | 343 |
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
|
344 |
``styles'': |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
345 |
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
346 |
\begin{quote}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
347 |
\verb!@!\verb!{thm_style stylename thm}!\\
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
348 |
\verb!@!\verb!{term_style stylename term}!
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
349 |
\end{quote}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
350 |
|
| 16040 | 351 |
A ``style'' is a transformation of propositions. There are predefined |
| 21558 | 352 |
styles, namely \verb!lhs! and \verb!rhs!, \verb!prem1! up to \verb!prem9!, and \verb!concl!. |
| 16166 | 353 |
For example, |
| 16076 | 354 |
the output |
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
355 |
\begin{center}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
356 |
\begin{tabular}{l@ {~~@{text "="}~~}l}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
357 |
@{thm_style lhs foldl_Nil} & @{thm_style rhs foldl_Nil}\\
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
358 |
@{thm_style lhs foldl_Cons} & @{thm_style rhs foldl_Cons}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
359 |
\end{tabular}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
360 |
\end{center}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
361 |
is produced by the following code: |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
362 |
\begin{quote}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
363 |
\verb!\begin{center}!\\
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
364 |
\verb!\begin{tabular}{l@ {~~!\verb!@!\verb!{text "="}~~}l}!\\
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
365 |
\verb!@!\verb!{thm_style lhs foldl_Nil} & @!\verb!{thm_style rhs foldl_Nil}!\\
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
366 |
\verb!@!\verb!{thm_style lhs foldl_Cons} & @!\verb!{thm_style rhs foldl_Cons}!\\
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
367 |
\verb!\end{tabular}!\\
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
368 |
\verb!\end{center}!
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
369 |
\end{quote}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
370 |
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
|
371 |
It prevents Isabelle from interpreting \verb!@ {~~...~~}!
|
| 16076 | 372 |
as an antiquotation. The styles \verb!lhs! and \verb!rhs! |
| 16040 | 373 |
extract the left hand side (or right hand side respectivly) from the |
| 16076 | 374 |
conclusion of propositions consisting of a binary operator |
| 16040 | 375 |
(e.~g.~@{text "="}, @{text "\<equiv>"}, @{text "<"}).
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
376 |
|
| 16165 | 377 |
Likewise, \verb!concl! may be used as a style to show just the |
| 16076 | 378 |
conclusion of a proposition. For example, take \verb!hd_Cons_tl!: |
| 16040 | 379 |
\begin{center}
|
|
17031
ffa73448025e
added hint for position of aqu options in connection with styles
haftmann
parents:
16395
diff
changeset
|
380 |
@{thm [show_types] hd_Cons_tl}
|
| 16040 | 381 |
\end{center}
|
382 |
To print just the conclusion, |
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
383 |
\begin{center}
|
|
17031
ffa73448025e
added hint for position of aqu options in connection with styles
haftmann
parents:
16395
diff
changeset
|
384 |
@{thm_style [show_types] concl hd_Cons_tl}
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
385 |
\end{center}
|
| 16040 | 386 |
type |
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
387 |
\begin{quote}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
388 |
\verb!\begin{center}!\\
|
|
17031
ffa73448025e
added hint for position of aqu options in connection with styles
haftmann
parents:
16395
diff
changeset
|
389 |
\verb!@!\verb!{thm_style [show_types] concl hd_Cons_tl}!\\
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
390 |
\verb!\end{center}!
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
391 |
\end{quote}
|
| 17127 | 392 |
Beware that any options must be placed \emph{before}
|
|
17031
ffa73448025e
added hint for position of aqu options in connection with styles
haftmann
parents:
16395
diff
changeset
|
393 |
the name of the style, as in this example. |
|
ffa73448025e
added hint for position of aqu options in connection with styles
haftmann
parents:
16395
diff
changeset
|
394 |
|
| 16166 | 395 |
Further use cases can be found in \S\ref{sec:yourself}.
|
396 |
||
| 15953 | 397 |
If you are not afraid of ML, you may also define your own styles. |
| 15960 | 398 |
A style is implemented by an ML function of type |
399 |
\verb!Proof.context -> term -> term!. |
|
400 |
Have a look at the following example: |
|
401 |
||
| 16040 | 402 |
*} |
| 16075 | 403 |
(*<*) |
| 16040 | 404 |
setup {*
|
405 |
let |
|
406 |
fun my_concl ctxt = Logic.strip_imp_concl |
|
| 18708 | 407 |
in TermStyle.add_style "my_concl" my_concl |
| 16040 | 408 |
end; |
409 |
*} |
|
| 16075 | 410 |
(*>*) |
| 16040 | 411 |
text {*
|
412 |
||
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
413 |
\begin{quote}
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
414 |
\verb!setup {!\verb!*!\\
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
415 |
\verb!let!\\ |
| 16040 | 416 |
\verb! fun my_concl ctxt = Logic.strip_imp_concl!\\ |
| 18708 | 417 |
\verb! in TermStyle.add_style "my_concl" my_concl!\\ |
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
418 |
\verb!end;!\\ |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
419 |
\verb!*!\verb!}!\\ |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
420 |
\end{quote}
|
| 15960 | 421 |
|
| 16076 | 422 |
\noindent |
| 16165 | 423 |
This example shows how the \verb!concl! style is implemented |
| 16076 | 424 |
and may be used as as a ``copy-and-paste'' pattern to write your own styles. |
| 15960 | 425 |
|
| 16076 | 426 |
The code should go into your theory file, separate from the \LaTeX\ text. |
| 16040 | 427 |
The \verb!let! expression avoids polluting the |
| 15960 | 428 |
ML global namespace. Each style receives the current proof context |
| 16076 | 429 |
as first argument; this is helpful in situations where the |
430 |
style has some object-logic specific behaviour for example. |
|
| 15960 | 431 |
|
432 |
The mapping from identifier name to the style function |
|
| 17123 | 433 |
is done by the @{ML TermStyle.add_style} expression which expects the desired
|
| 15960 | 434 |
style name and the style function as arguments. |
435 |
||
436 |
After this \verb!setup!, |
|
| 16040 | 437 |
there will be a new style available named \verb!my_concl!, thus allowing |
438 |
antiquoations like \verb!@!\verb!{thm_style my_concl hd_Cons_tl}!
|
|
439 |
yielding @{thm_style my_concl hd_Cons_tl}.
|
|
|
15917
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
440 |
|
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
441 |
*} |
|
cd4983c76548
Added short description of thm_style and term_style antiquotation
haftmann
parents:
15689
diff
changeset
|
442 |
|
| 15337 | 443 |
(*<*) |
444 |
end |
|
| 16175 | 445 |
(*>*) |