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