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