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