doc-src/LaTeXsugar/Sugar/document/Sugar.tex
author wenzelm
Tue, 10 Mar 2009 11:01:15 +0100
changeset 30405 bd38973f39e5
parent 29297 62e0f892e525
child 30502 b80d2621caee
permissions -rw-r--r--
updated generated file -- changed due to different treatmeant of type constraints in OptionalSugar.thy;

%
\begin{isabellebody}%
\def\isabellecontext{Sugar}%
%
\isadelimtheory
%
\endisadelimtheory
%
\isatagtheory
%
\endisatagtheory
{\isafoldtheory}%
%
\isadelimtheory
%
\endisadelimtheory
%
\isamarkupsection{Introduction%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
This document is for those Isabelle users who have mastered
the art of mixing \LaTeX\ text and Isabelle theories and never want to
typeset a theorem by hand anymore because they have experienced the
bliss of writing \verb!@!\verb!{thm[display]setsum_cartesian_product[no_vars]}!
and seeing Isabelle typeset it for them:
\begin{isabelle}%
{\isacharparenleft}{\isasymSum}x{\isasymin}A{\isachardot}\ {\isasymSum}y{\isasymin}B{\isachardot}\ f\ x\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymSum}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isasymin}A\ {\isasymtimes}\ B{\isachardot}\ f\ x\ y{\isacharparenright}%
\end{isabelle}
No typos, no omissions, no sweat.
If you have not experienced that joy, read Chapter 4, \emph{Presenting
Theories}, \cite{LNCS2283} first.

If you have mastered the art of Isabelle's \emph{antiquotations},
i.e.\ things like the above \verb!@!\verb!{thm...}!, beware: in your vanity
you may be tempted to think that all readers of the stunning ps or pdf
documents you can now produce at the drop of a hat will be struck with
awe at the beauty unfolding in front of their eyes. Until one day you
come across that very critical of readers known as the ``common referee''.
He has the nasty habit of refusing to understand unfamiliar notation
like Isabelle's infamous \isa{{\isasymlbrakk}\ {\isasymrbrakk}\ {\isasymLongrightarrow}} no matter how many times you
explain it in your paper. Even worse, he thinks that using \isa{{\isasymlbrakk}\ {\isasymrbrakk}} for anything other than denotational semantics is a cardinal sin
that must be punished by instant rejection.


This document shows you how to make Isabelle and \LaTeX\ cooperate to
produce ordinary looking mathematics that hides the fact that it was
typeset by a machine. You merely need to load the right files:
\begin{itemize}
\item Import theory \texttt{LaTeXsugar} in the header of your own
theory.  You may also want bits of \texttt{OptionalSugar}, which you can
copy selectively into your own theory or import as a whole.  Both
theories live in \texttt{HOL/Library} and are found automatically.

\item Should you need additional \LaTeX\ packages (the text will tell
you so), you include them at the beginning of your \LaTeX\ document,
typically in \texttt{root.tex}. For a start, you should
\verb!\usepackage{amssymb}! --- otherwise typesetting
\isa{{\isachardoublequote}{\isasymnot}{\isacharparenleft}{\isasymexists}x{\isachardot}\ P\ x{\isacharparenright}{\isachardoublequote}} will fail because the AMS symbol
\isa{{\isasymnexists}} is missing.
\end{itemize}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsection{HOL syntax%
}
\isamarkuptrue%
%
\isamarkupsubsection{Logic%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The formula \isa{{\isachardoublequote}{\isasymnot}{\isacharparenleft}{\isasymexists}x{\isachardot}\ P\ x{\isacharparenright}{\isachardoublequote}} is typeset as \isa{{\isasymnexists}x{\isachardot}\ P\ x}.

The predefined constructs \isa{if}, \isa{let} and
\isa{case} are set in sans serif font to distinguish them from
other functions. This improves readability:
\begin{itemize}
\item \isa{\textsf{if}\ b\ \textsf{then}\ e\isactrlisub {\isadigit{1}}\ \textsf{else}\ e\isactrlisub {\isadigit{2}}} instead of \isa{if\ b\ then\ e\isactrlisub {\isadigit{1}}\ else\ e\isactrlisub {\isadigit{2}}}.
\item \isa{\textsf{let}\ x\ {\isacharequal}\ e\isactrlisub {\isadigit{1}}\ \textsf{in}\ e\isactrlisub {\isadigit{2}}} instead of \isa{let\ x\ {\isacharequal}\ e\isactrlisub {\isadigit{1}}\ in\ e\isactrlisub {\isadigit{2}}}.
\item \isa{\textsf{case}\ x\ \textsf{of}\ True\ {\isasymRightarrow}\ e\isactrlisub {\isadigit{1}}\ {\isacharbar}\ False\ {\isasymRightarrow}\ e\isactrlisub {\isadigit{2}}} instead of\\
      \isa{case\ x\ of\ True\ {\isasymRightarrow}\ e\isactrlisub {\isadigit{1}}\ {\isacharbar}\ False\ {\isasymRightarrow}\ e\isactrlisub {\isadigit{2}}}.
\end{itemize}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Sets%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
Although set syntax in HOL is already close to
standard, we provide a few further improvements:
\begin{itemize}
\item \isa{{\isacharbraceleft}x\ {\isacharbar}\ P{\isacharbraceright}} instead of \isa{{\isacharbraceleft}x{\isachardot}\ P{\isacharbraceright}}.
\item \isa{{\isasymemptyset}} instead of \isa{{\isacharbraceleft}{\isacharbraceright}}, where
 \isa{{\isasymemptyset}} is also input syntax.
\item \isa{{\isacharbraceleft}a{\isacharcomma}\ b{\isacharcomma}\ c{\isacharbraceright}\ {\isasymunion}\ M} instead of \isa{insert\ a\ {\isacharparenleft}insert\ b\ {\isacharparenleft}insert\ c\ M{\isacharparenright}{\isacharparenright}}.
\end{itemize}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Lists%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
If lists are used heavily, the following notations increase readability:
\begin{itemize}
\item \isa{x{\isasymcdot}xs} instead of \isa{x\ {\isacharhash}\ xs},
      where \isa{x{\isasymcdot}xs} is also input syntax.
If you prefer more space around the $\cdot$ you have to redefine
\verb!\isasymcdot! in \LaTeX:
\verb!\renewcommand{\isasymcdot}{\isamath{\,\cdot\,}}!

\item \isa{{\isacharbar}xs{\isacharbar}} instead of \isa{length\ xs}.
\item \isa{xs\ensuremath{_{[\mathit{n}]}}} instead of \isa{nth\ xs\ n},
      the $n$th element of \isa{xs}.

\item Human readers are good at converting automatically from lists to
sets. Hence \texttt{OptionalSugar} contains syntax for supressing the
conversion function \isa{set}: for example, \isa{{\isachardoublequote}x\ {\isasymin}\ set\ xs{\isachardoublequote}}
becomes \isa{x\ {\isasymin}\ xs}.

\item The \isa{{\isacharat}} operation associates implicitly to the right,
which leads to unpleasant line breaks if the term is too long for one
line. To avoid this, \texttt{OptionalSugar} contains syntax to group
\isa{{\isacharat}}-terms to the left before printing, which leads to better
line breaking behaviour:
\begin{isabelle}%
term\isactrlisub {\isadigit{0}}\ \isacharat\ term\isactrlisub {\isadigit{1}}\ \isacharat\ term\isactrlisub {\isadigit{2}}\ \isacharat\ term\isactrlisub {\isadigit{3}}\ \isacharat\ term\isactrlisub {\isadigit{4}}\ \isacharat\ term\isactrlisub {\isadigit{5}}\ \isacharat\ term\isactrlisub {\isadigit{6}}\ \isacharat\ term\isactrlisub {\isadigit{7}}\ \isacharat\ term\isactrlisub {\isadigit{8}}\ \isacharat\ term\isactrlisub {\isadigit{9}}\ \isacharat\ term\isactrlisub {\isadigit{1}}\isactrlisub {\isadigit{0}}%
\end{isabelle}

\end{itemize}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsection{Printing theorems%
}
\isamarkuptrue%
%
\isamarkupsubsection{Question marks%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
If you print anything, especially theorems, containing
schematic variables they are prefixed with a question mark:
\verb!@!\verb!{thm conjI}! results in \isa{{\isasymlbrakk}P{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isasymand}\ Q}. Most of the time
you would rather not see the question marks. There is an attribute
\verb!no_vars! that you can attach to the theorem that turns its
schematic into ordinary free variables: \verb!@!\verb!{thm conjI[no_vars]}!
results in \isa{{\isasymlbrakk}P{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isasymand}\ Q}.

This \verb!no_vars! business can become a bit tedious.
If you would rather never see question marks, simply put
\begin{verbatim}
reset show_question_marks;
\end{verbatim}
at the beginning of your file \texttt{ROOT.ML}.
The rest of this document is produced with this flag reset.

Hint: Resetting \verb!show_question_marks! only supresses question
marks; variables that end in digits, e.g. \isa{x{\isadigit{1}}}, are still
printed with a trailing \isa{{\isachardot}{\isadigit{0}}}, e.g. \isa{x{\isadigit{1}}{\isachardot}{\isadigit{0}}}, their
internal index. This can be avoided by turning the last digit into a
subscript: write \verb!x\<^isub>1! and obtain the much nicer \isa{x\isactrlisub {\isadigit{1}}}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isadelimML
%
\endisadelimML
%
\isatagML
%
\endisatagML
{\isafoldML}%
%
\isadelimML
%
\endisadelimML
%
\isamarkupsubsection{Qualified names%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
If there are multiple declarations of the same name, Isabelle prints
the qualified name, for example \isa{T{\isachardot}length}, where \isa{T} is the
theory it is defined in, to distinguish it from the predefined \isa{{\isachardoublequote}List{\isachardot}length{\isachardoublequote}}. In case there is no danger of confusion, you can insist on
short names (no qualifiers) by setting \verb!short_names!, typically
in \texttt{ROOT.ML}:
\begin{verbatim}
set short_names;
\end{verbatim}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Variable names\label{sec:varnames}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
It sometimes happens that you want to change the name of a
variable in a theorem before printing it. This can easily be achieved
with the help of Isabelle's instantiation attribute \texttt{where}:
\isa{{\isasymlbrakk}{\isasymphi}{\isacharsemicolon}\ {\isasympsi}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymphi}\ {\isasymand}\ {\isasympsi}} is the result of
\begin{quote}
\verb!@!\verb!{thm conjI[where P = \<phi> and Q = \<psi>]}!
\end{quote}
To support the ``\_''-notation for irrelevant variables
the constant \texttt{DUMMY} has been introduced:
\isa{fst\ {\isacharparenleft}a{\isacharcomma}\ \_{\isacharparenright}\ {\isacharequal}\ a} is produced by
\begin{quote}
\verb!@!\verb!{thm fst_conv[where b = DUMMY]}!
\end{quote}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Inference rules%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
To print theorems as inference rules you need to include Didier
R\'emy's \texttt{mathpartir} package~\cite{mathpartir}
for typesetting inference rules in your \LaTeX\ file.

Writing \verb!@!\verb!{thm[mode=Rule] conjI}! produces
\isa{\mbox{}\inferrule{\mbox{P}\\\ \mbox{Q}}{\mbox{P\ {\isasymand}\ Q}}}, even in the middle of a sentence.
If you prefer your inference rule on a separate line, maybe with a name,
\begin{center}
\isa{\mbox{}\inferrule{\mbox{P}\\\ \mbox{Q}}{\mbox{P\ {\isasymand}\ Q}}} {\sc conjI}
\end{center}
is produced by
\begin{quote}
\verb!\begin{center}!\\
\verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI}!\\
\verb!\end{center}!
\end{quote}
It is not recommended to use the standard \texttt{display} option
together with \texttt{Rule} because centering does not work and because
the line breaking mechanisms of \texttt{display} and \texttt{mathpartir} can
clash.

Of course you can display multiple rules in this fashion:
\begin{quote}
\verb!\begin{center}!\\
\verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI} \\[1ex]!\\
\verb!@!\verb!{thm[mode=Rule] conjE} {\sc disjI$_1$} \qquad!\\
\verb!@!\verb!{thm[mode=Rule] disjE} {\sc disjI$_2$}!\\
\verb!\end{center}!
\end{quote}
yields
\begin{center}\small
\isa{\mbox{}\inferrule{\mbox{P}\\\ \mbox{Q}}{\mbox{P\ {\isasymand}\ Q}}} {\sc conjI} \\[1ex]
\isa{\mbox{}\inferrule{\mbox{P}}{\mbox{P\ {\isasymor}\ Q}}} {\sc disjI$_1$} \qquad
\isa{\mbox{}\inferrule{\mbox{Q}}{\mbox{P\ {\isasymor}\ Q}}} {\sc disjI$_2$}
\end{center}

The \texttt{mathpartir} package copes well if there are too many
premises for one line:
\begin{center}
\isa{\mbox{}\inferrule{\mbox{A\ {\isasymlongrightarrow}\ B}\\\ \mbox{B\ {\isasymlongrightarrow}\ C}\\\ \mbox{C\ {\isasymlongrightarrow}\ D}\\\ \mbox{D\ {\isasymlongrightarrow}\ E}\\\ \mbox{E\ {\isasymlongrightarrow}\ F}\\\ \mbox{F\ {\isasymlongrightarrow}\ G}\\\ \mbox{G\ {\isasymlongrightarrow}\ H}\\\ \mbox{H\ {\isasymlongrightarrow}\ I}\\\ \mbox{I\ {\isasymlongrightarrow}\ J}\\\ \mbox{J\ {\isasymlongrightarrow}\ K}}{\mbox{A\ {\isasymlongrightarrow}\ K}}}
\end{center}

Limitations: 1. Premises and conclusion must each not be longer than
the line.  2. Premises that are \isa{{\isasymLongrightarrow}}-implications are again
displayed with a horizontal line, which looks at least unusual.


In case you print theorems without premises no rule will be printed by the
\texttt{Rule} print mode. However, you can use \texttt{Axiom} instead:
\begin{quote}
\verb!\begin{center}!\\
\verb!@!\verb!{thm[mode=Axiom] refl} {\sc refl}! \\
\verb!\end{center}!
\end{quote}
yields
\begin{center}
\isa{\mbox{}\inferrule{\mbox{}}{\mbox{t\ {\isacharequal}\ t}}} {\sc refl} 
\end{center}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Displays and font sizes%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
When displaying theorems with the \texttt{display} option, e.g.
\verb!@!\verb!{thm[display] refl}! \begin{isabelle}%
t\ {\isacharequal}\ t%
\end{isabelle} the theorem is
set in small font. It uses the \LaTeX-macro \verb!\isastyle!,
which is also the style that regular theory text is set in, e.g.%
\end{isamarkuptext}%
\isamarkuptrue%
\isacommand{lemma}\isamarkupfalse%
\ {\isachardoublequoteopen}t\ {\isacharequal}\ t{\isachardoublequoteclose}%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\begin{isamarkuptext}%
\noindent Otherwise \verb!\isastyleminor! is used,
which does not modify the font size (assuming you stick to the default
\verb!\isabellestyle{it}! in \texttt{root.tex}). If you prefer
normal font size throughout your text, include
\begin{quote}
\verb!\renewcommand{\isastyle}{\isastyleminor}!
\end{quote}
in \texttt{root.tex}. On the other hand, if you like the small font,
just put \verb!\isastyle! in front of the text in question,
e.g.\ at the start of one of the center-environments above.

The advantage of the display option is that you can display a whole
list of theorems in one go. For example,
\verb!@!\verb!{thm[display] foldl.simps}!
generates \begin{isabelle}%
foldl\ f\ a\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ a\isasep\isanewline%
foldl\ f\ a\ {\isacharparenleft}x{\isasymcdot}xs{\isacharparenright}\ {\isacharequal}\ foldl\ f\ {\isacharparenleft}f\ a\ x{\isacharparenright}\ xs%
\end{isabelle}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{If-then%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
If you prefer a fake ``natural language'' style you can produce
the body of
\newtheorem{theorem}{Theorem}
\begin{theorem}
\isa{{\normalsize{}If\,}\ \mbox{i\ {\isasymle}\ j}\ {\normalsize \,and\,}\ \mbox{j\ {\isasymle}\ k}\ {\normalsize \,then\,}\ i\ {\isasymle}\ k{\isachardot}}
\end{theorem}
by typing
\begin{quote}
\verb!@!\verb!{thm[mode=IfThen] le_trans}!
\end{quote}

In order to prevent odd line breaks, the premises are put into boxes.
At times this is too drastic:
\begin{theorem}
\isa{{\normalsize{}If\,}\ \mbox{longpremise}\ {\normalsize \,and\,}\ \mbox{longerpremise}\ {\normalsize \,and\,}\ \mbox{P\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}}\ {\normalsize \,and\,}\ \mbox{longestpremise}\ {\normalsize \,then\,}\ conclusion{\isachardot}}
\end{theorem}
In which case you should use \texttt{IfThenNoBox} instead of
\texttt{IfThen}:
\begin{theorem}
\isa{{\normalsize{}If\,}\ longpremise\ {\normalsize \,and\,}\ longerpremise\ {\normalsize \,and\,}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}\ {\normalsize \,and\,}\ longestpremise\ {\normalsize \,then\,}\ conclusion{\isachardot}}
\end{theorem}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Doing it yourself\label{sec:yourself}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
If for some reason you want or need to present theorems your
own way, you can extract the premises and the conclusion explicitly
and combine them as you like:
\begin{itemize}
\item \verb!@!\verb!{thm_style prem1! $thm$\verb!}!
prints premise 1 of $thm$ (and similarly up to \texttt{prem9}).
\item \verb!@!\verb!{thm_style concl! $thm$\verb!}!
prints the conclusion of $thm$.
\end{itemize}
For example, ``from \isa{Q} and
\isa{P} we conclude \isa{P\ {\isasymand}\ Q}''
is produced by
\begin{quote}
\verb!from !\verb!@!\verb!{thm_style prem2 conjI}! \verb!and !\verb!@!\verb!{thm_style prem1 conjI}!\\
\verb!we conclude !\verb!@!\verb!{thm_style concl conjI}!
\end{quote}
Thus you can rearrange or hide premises and typeset the theorem as you like.
The \verb!thm_style! antiquotation is a general mechanism explained
in \S\ref{sec:styles}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsubsection{Patterns%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
In \S\ref{sec:varnames} we shows how to create patterns containing
  ``\isa{\_}''.
  You can drive this game even further and extend the syntax of let
  bindings such that certain functions like \isa{fst}, \isa{hd}, 
  etc.\ are printed as patterns. \texttt{OptionalSugar} provides the
  following:
  
  \begin{center}
  \begin{tabular}{l@ {~~produced by~~}l}
  \isa{\textsf{let}\ {\isacharparenleft}x{\isacharcomma}\ \_{\isacharparenright}\ {\isacharequal}\ p\ \textsf{in}\ t} & \verb!@!\verb!{term "let x = fst p in t"}!\\
  \isa{\textsf{let}\ {\isacharparenleft}\_{\isacharcomma}\ x{\isacharparenright}\ {\isacharequal}\ p\ \textsf{in}\ t} & \verb!@!\verb!{term "let x = snd p in t"}!\\
  \isa{\textsf{let}\ x{\isasymcdot}\_\ {\isacharequal}\ xs\ \textsf{in}\ t} & \verb!@!\verb!{term "let x = hd xs in t"}!\\
  \isa{\textsf{let}\ \_{\isasymcdot}x\ {\isacharequal}\ xs\ \textsf{in}\ t} & \verb!@!\verb!{term "let x = tl xs in t"}!\\
  \isa{\textsf{let}\ Some\ x\ {\isacharequal}\ y\ \textsf{in}\ t} & \verb!@!\verb!{term "let x = the y in t"}!\\
  \end{tabular}
  \end{center}%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsection{Proofs%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
Full proofs, even if written in beautiful Isar style, are
likely to be too long and detailed to be included in conference
papers, but some key lemmas might be of interest.
It is usually easiest to put them in figures like the one in Fig.\
\ref{fig:proof}. This was achieved with the \isakeyword{text\_raw} command:%
\end{isamarkuptext}%
\isamarkuptrue%
%
\begin{figure}
  \begin{center}\begin{minipage}{0.6\textwidth}  
  \isastyleminor\isamarkuptrue
\isacommand{lemma}\isamarkupfalse%
\ True\isanewline
%
\isadelimproof
%
\endisadelimproof
%
\isatagproof
\isacommand{proof}\isamarkupfalse%
\ {\isacharminus}\isanewline
\ \ %
\isamarkupcmt{pretty trivial%
}
\isanewline
\ \ \isacommand{show}\isamarkupfalse%
\ True\ \isacommand{by}\isamarkupfalse%
\ force\isanewline
\isacommand{qed}\isamarkupfalse%
%
\endisatagproof
{\isafoldproof}%
%
\isadelimproof
%
\endisadelimproof
%
\end{minipage}\end{center}
  \caption{Example proof in a figure.}\label{fig:proof}
  \end{figure}
%
\begin{isamarkuptext}%
\begin{quote}
\small
\verb!text_raw {!\verb!*!\\
\verb!  \begin{figure}!\\
\verb!  \begin{center}\begin{minipage}{0.6\textwidth}!\\
\verb!  \isastyleminor\isamarkuptrue!\\
\verb!*!\verb!}!\\
\verb!lemma True!\\
\verb!proof -!\\
\verb!  -- "pretty trivial"!\\
\verb!  show True by force!\\
\verb!qed!\\
\verb!text_raw {!\verb!*!\\
\verb!  \end{minipage}\end{center}!\\
\verb!  \caption{Example proof in a figure.}\label{fig:proof}!\\
\verb!  \end{figure}!\\
\verb!*!\verb!}!
\end{quote}

Other theory text, e.g.\ definitions, can be put in figures, too.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isamarkupsection{Styles\label{sec:styles}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
The \verb!thm! antiquotation works nicely for single theorems, but
  sets of equations as used in definitions are more difficult to
  typeset nicely: people tend to prefer aligned \isa{{\isacharequal}} signs.

  To deal with such cases where it is desirable to dive into the structure
  of terms and theorems, Isabelle offers antiquotations featuring
  ``styles'':

    \begin{quote}
    \verb!@!\verb!{thm_style stylename thm}!\\
    \verb!@!\verb!{term_style stylename term}!
    \end{quote}

  A ``style'' is a transformation of propositions. There are predefined
  styles, namely \verb!lhs! and \verb!rhs!, \verb!prem1! up to \verb!prem9!, and \verb!concl!.
  For example, 
  the output
  \begin{center}
  \begin{tabular}{l@ {~~\isa{{\isacharequal}}~~}l}
  \isa{foldl\ f\ a\ {\isacharbrackleft}{\isacharbrackright}} & \isa{a}\\
  \isa{foldl\ f\ a\ {\isacharparenleft}x{\isasymcdot}xs{\isacharparenright}} & \isa{foldl\ f\ {\isacharparenleft}f\ a\ x{\isacharparenright}\ xs}
  \end{tabular}
  \end{center}
  is produced by the following code:
  \begin{quote}
    \verb!\begin{center}!\\
    \verb!\begin{tabular}{l@ {~~!\verb!@!\verb!{text "="}~~}l}!\\
    \verb!@!\verb!{thm_style lhs foldl_Nil} & @!\verb!{thm_style rhs foldl_Nil}!\\
    \verb!@!\verb!{thm_style lhs foldl_Cons} & @!\verb!{thm_style rhs foldl_Cons}!\\
    \verb!\end{tabular}!\\
    \verb!\end{center}!
  \end{quote}
  Note the space between \verb!@! and \verb!{! in the tabular argument.
  It prevents Isabelle from interpreting \verb!@ {~~...~~}! 
  as an antiquotation. The styles \verb!lhs! and \verb!rhs!
  extract the left hand side (or right hand side respectively) from the
  conclusion of propositions consisting of a binary operator
  (e.~g.~\isa{{\isacharequal}}, \isa{{\isasymequiv}}, \isa{{\isacharless}}).

  Likewise, \verb!concl! may be used as a style to show just the
  conclusion of a proposition. For example, take \verb!hd_Cons_tl!:
  \begin{center}
    \isa{{\isacharparenleft}xs\ {\isasymColon}\ {\isacharprime}a\ list{\isacharparenright}\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd\ xs{\isasymcdot}tl\ xs\ {\isacharequal}\ xs}
  \end{center}
  To print just the conclusion,
  \begin{center}
    \isa{hd\ {\isacharparenleft}xs\ {\isasymColon}\ {\isacharprime}a\ list{\isacharparenright}{\isasymcdot}tl\ xs\ {\isacharequal}\ xs}
  \end{center}
  type
  \begin{quote}
    \verb!\begin{center}!\\
    \verb!@!\verb!{thm_style [show_types] concl hd_Cons_tl}!\\
    \verb!\end{center}!
  \end{quote}
  Beware that any options must be placed \emph{before}
  the name of the style, as in this example.

  Further use cases can be found in \S\ref{sec:yourself}.

  If you are not afraid of ML, you may also define your own styles.
  A style is implemented by an ML function of type
  \verb!Proof.context -> term -> term!.
  Have a look at the following example:%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isadelimML
%
\endisadelimML
%
\isatagML
%
\endisatagML
{\isafoldML}%
%
\isadelimML
%
\endisadelimML
%
\begin{isamarkuptext}%
\begin{quote}
    \verb!setup {!\verb!*!\\
    \verb!let!\\
    \verb!  fun my_concl ctxt = Logic.strip_imp_concl!\\
    \verb!  in TermStyle.add_style "my_concl" my_concl!\\
    \verb!end;!\\
    \verb!*!\verb!}!\\
  \end{quote}

  \noindent
  This example shows how the \verb!concl! style is implemented
  and may be used as as a ``copy-and-paste'' pattern to write your own styles.

  The code should go into your theory file, separate from the \LaTeX\ text.
  The \verb!let! expression avoids polluting the
  ML global namespace. Each style receives the current proof context
  as first argument; this is helpful in situations where the
  style has some object-logic specific behaviour for example.

  The mapping from identifier name to the style function
  is done by the \verb|TermStyle.add_style| expression which expects the desired
  style name and the style function as arguments.
  
  After this \verb!setup!,
  there will be a new style available named \verb!my_concl!, thus allowing
  antiquoations like \verb!@!\verb!{thm_style my_concl hd_Cons_tl}!
  yielding \isa{hd\ xs{\isasymcdot}tl\ xs\ {\isacharequal}\ xs}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
\isadelimtheory
%
\endisadelimtheory
%
\isatagtheory
%
\endisatagtheory
{\isafoldtheory}%
%
\isadelimtheory
%
\endisadelimtheory
\end{isabellebody}%
%%% Local Variables:
%%% mode: latex
%%% TeX-master: "root"
%%% End: