more standard document preparation within session context: avoid clashes with generated .tex files, even on case-insensible file-system;
--- a/doc-src/ROOT Tue Aug 28 13:15:15 2012 +0200
+++ b/doc-src/ROOT Tue Aug 28 14:37:57 2012 +0200
@@ -278,9 +278,7 @@
"document/root.tex"
session Tutorial (doc) in "TutorialI" = HOL +
- options [browser_info = false, document = false,
- document_dump = document, document_dump_mode = "tex",
- print_mode = "brackets"]
+ options [document_variants = "tutorial", print_mode = "brackets"]
theories [thy_output_indent = 5]
"ToyList/ToyList"
"Ifexpr/Ifexpr"
@@ -317,7 +315,7 @@
theories
"Protocol/NS_Public"
"Documents/Documents"
- theories [document_dump = ""]
+ theories [document = ""]
"Types/Setup"
theories [pretty_margin = 64, thy_output_indent = 0]
"Types/Numbers"
@@ -338,8 +336,15 @@
"Sets/Functions"
"Sets/Relations"
"Sets/Recur"
+ files
+ "ToyList/ToyList1"
+ "ToyList/ToyList2"
+ "../pdfsetup.sty"
+ "../proof.sty"
+ "../ttbox.sty"
+ "../manual.bib"
+ "document/pghead.eps"
+ "document/pghead.pdf"
+ "document/build"
+ "document/root.tex"
-session "HOL-ToyList2" (doc) in "TutorialI/ToyList2" = HOL +
- options [browser_info = false, document = false, print_mode = "brackets"]
- theories ToyList
-
--- a/doc-src/TutorialI/Advanced/advanced.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,49 +0,0 @@
-\chapter{Advanced Simplification and Induction}
-
-Although we have already learned a lot about simplification and
-induction, there are some advanced proof techniques that we have not covered
-yet and which are worth learning. The sections of this chapter are
-independent of each other and can be read in any order.
-
-\input{document/simp2.tex}
-
-\section{Advanced Induction Techniques}
-\label{sec:advanced-ind}
-\index{induction|(}
-\input{document/AdvancedInd.tex}
-\input{document/CTLind.tex}
-\index{induction|)}
-
-%\section{Advanced Forms of Recursion}
-%\index{recdef@\isacommand {recdef} (command)|(}
-
-%This section introduces advanced forms of
-%\isacommand{recdef}: how to establish termination by means other than measure
-%functions, how to define recursive functions over nested recursive datatypes
-%and how to deal with partial functions.
-%
-%If, after reading this section, you feel that the definition of recursive
-%functions is overly complicated by the requirement of
-%totality, you should ponder the alternatives. In a logic of partial functions,
-%recursive definitions are always accepted. But there are many
-%such logics, and no clear winner has emerged. And in all of these logics you
-%are (more or less frequently) required to reason about the definedness of
-%terms explicitly. Thus one shifts definedness arguments from definition time to
-%proof time. In HOL you may have to work hard to define a function, but proofs
-%can then proceed unencumbered by worries about undefinedness.
-
-%\subsection{Beyond Measure}
-%\label{sec:beyond-measure}
-%\input{document/WFrec.tex}
-%
-%\subsection{Recursion Over Nested Datatypes}
-%\label{sec:nested-recdef}
-%\input{document/Nested0.tex}
-%\input{document/Nested1.tex}
-%\input{document/Nested2.tex}
-%
-%\subsection{Partial Functions}
-%\index{functions!partial}
-%\input{document/Partial.tex}
-%
-%\index{recdef@\isacommand {recdef} (command)|)}
--- a/doc-src/TutorialI/CTL/ctl.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,6 +0,0 @@
-\index{model checking example|(}%
-\index{lfp@{\texttt{lfp}}!applications of|see{CTL}}
-\input{document/Base.tex}
-\input{document/PDL.tex}
-\input{document/CTL.tex}
-\index{model checking example|)}
--- a/doc-src/TutorialI/Datatype/Nested.thy Tue Aug 28 13:15:15 2012 +0200
+++ b/doc-src/TutorialI/Datatype/Nested.thy Tue Aug 28 14:37:57 2012 +0200
@@ -30,7 +30,7 @@
would be something like
\medskip
-\input{document/unfoldnested.tex}
+\input{unfoldnested.tex}
\medskip
\noindent
--- a/doc-src/TutorialI/Documents/documents.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-
-\chapter{Presenting Theories}
-\label{ch:thy-present}
-
-By now the reader should have become sufficiently acquainted with elementary
-theory development in Isabelle/HOL\@. The following interlude describes
-how to present theories in a typographically
-pleasing manner. Isabelle provides a rich infrastructure for concrete syntax
-of the underlying $\lambda$-calculus language (see
-{\S}\ref{sec:concrete-syntax}), as well as document preparation of theory texts
-based on existing PDF-{\LaTeX} technology (see
-{\S}\ref{sec:document-preparation}).
-
-As pointed out by Leibniz\index{Leibniz, Gottfried Wilhelm} more than 300
-years ago, \emph{notions} are in principle more important than
-\emph{notations}, but suggestive textual representation of ideas is vital to
-reduce the mental effort to comprehend and apply them.
-
-\input{document/Documents.tex}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
--- a/doc-src/TutorialI/Inductive/inductive.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,31 +0,0 @@
-\chapter{Inductively Defined Sets} \label{chap:inductive}
-\index{inductive definitions|(}
-
-This chapter is dedicated to the most important definition principle after
-recursive functions and datatypes: inductively defined sets.
-
-We start with a simple example: the set of even numbers. A slightly more
-complicated example, the reflexive transitive closure, is the subject of
-{\S}\ref{sec:rtc}. In particular, some standard induction heuristics are
-discussed. Advanced forms of inductive definitions are discussed in
-{\S}\ref{sec:adv-ind-def}. To demonstrate the versatility of inductive
-definitions, the chapter closes with a case study from the realm of
-context-free grammars. The first two sections are required reading for anybody
-interested in mathematical modelling.
-
-\begin{warn}
-Predicates can also be defined inductively.
-See {\S}\ref{sec:ind-predicates}.
-\end{warn}
-
-\input{document/Even}
-\input{document/Mutual}
-\input{document/Star}
-
-\section{Advanced Inductive Definitions}
-\label{sec:adv-ind-def}
-\input{document/Advanced}
-
-\input{document/AB}
-
-\index{inductive definitions|)}
--- a/doc-src/TutorialI/Protocol/protocol.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,135 +0,0 @@
-\chapter{Case Study: Verifying a Security Protocol}
-\label{chap:crypto}
-
-\index{protocols!security|(}
-
-%crypto primitives
-\def\lbb{\mathopen{\{\kern-.30em|}}
-\def\rbb{\mathclose{|\kern-.32em\}}}
-\def\comp#1{\lbb#1\rbb}
-
-Communications security is an ancient art. Julius Caesar is said to have
-encrypted his messages, shifting each letter three places along the
-alphabet. Mary Queen of Scots was convicted of treason after a cipher used
-in her letters was broken. Today's postal system
-incorporates security features. The envelope provides a degree of
-\emph{secrecy}. The signature provides \emph{authenticity} (proof of
-origin), as do departmental stamps and letterheads.
-
-Networks are vulnerable: messages pass through many computers, any of which
-might be controlled by an adversary, who thus can capture or redirect
-messages. People who wish to communicate securely over such a network can
-use cryptography, but if they are to understand each other, they need to
-follow a
-\emph{protocol}: a pre-arranged sequence of message formats.
-
-Protocols can be attacked in many ways, even if encryption is unbreakable.
-A \emph{splicing attack} involves an adversary's sending a message composed
-of parts of several old messages. This fake message may have the correct
-format, fooling an honest party. The adversary might be able to masquerade
-as somebody else, or he might obtain a secret key.
-
-\emph{Nonces} help prevent splicing attacks. A typical nonce is a 20-byte
-random number. Each message that requires a reply incorporates a nonce. The
-reply must include a copy of that nonce, to prove that it is not a replay of
-a past message. The nonce in the reply must be cryptographically
-protected, since otherwise an adversary could easily replace it by a
-different one. You should be starting to see that protocol design is
-tricky!
-
-Researchers are developing methods for proving the correctness of security
-protocols. The Needham-Schroeder public-key
-protocol~\cite{needham-schroeder} has become a standard test case.
-Proposed in 1978, it was found to be defective nearly two decades
-later~\cite{lowe-fdr}. This toy protocol will be useful in demonstrating
-how to verify protocols using Isabelle.
-
-
-\section{The Needham-Schroeder Public-Key Protocol}\label{sec:ns-protocol}
-
-\index{Needham-Schroeder protocol|(}%
-This protocol uses public-key cryptography. Each person has a private key, known only to
-himself, and a public key, known to everybody. If Alice wants to send Bob a secret message, she
-encrypts it using Bob's public key (which everybody knows), and sends it to Bob. Only Bob has the
-matching private key, which is needed in order to decrypt Alice's message.
-
-The core of the Needham-Schroeder protocol consists of three messages:
-\begin{alignat*}{2}
- &1.&\quad A\to B &: \comp{Na,A}\sb{Kb} \\
- &2.&\quad B\to A &: \comp{Na,Nb}\sb{Ka} \\
- &3.&\quad A\to B &: \comp{Nb}\sb{Kb}
-\end{alignat*}
-First, let's understand the notation. In the first message, Alice
-sends Bob a message consisting of a nonce generated by Alice~($Na$)
-paired with Alice's name~($A$) and encrypted using Bob's public
-key~($Kb$). In the second message, Bob sends Alice a message
-consisting of $Na$ paired with a nonce generated by Bob~($Nb$),
-encrypted using Alice's public key~($Ka$). In the last message, Alice
-returns $Nb$ to Bob, encrypted using his public key.
-
-When Alice receives Message~2, she knows that Bob has acted on her
-message, since only he could have decrypted
-$\comp{Na,A}\sb{Kb}$ and extracted~$Na$. That is precisely what
-nonces are for. Similarly, message~3 assures Bob that Alice is
-active. But the protocol was widely believed~\cite{ban89} to satisfy a
-further property: that
-$Na$ and~$Nb$ were secrets shared by Alice and Bob. (Many
-protocols generate such shared secrets, which can be used
-to lessen the reliance on slow public-key operations.)
-Lowe\index{Lowe, Gavin|(} found this
-claim to be false: if Alice runs the protocol with someone untrustworthy
-(Charlie say), then he can start a new run with another agent (Bob say).
-Charlie uses Alice as an oracle, masquerading as
-Alice to Bob~\cite{lowe-fdr}.
-\begin{alignat*}{4}
- &1.&\quad A\to C &: \comp{Na,A}\sb{Kc} &&
- \qquad 1'.&\quad C\to B &: \comp{Na,A}\sb{Kb} \\
- &2.&\quad B\to A &: \comp{Na,Nb}\sb{Ka} \\
- &3.&\quad A\to C &: \comp{Nb}\sb{Kc} &&
- \qquad 3'.&\quad C\to B &: \comp{Nb}\sb{Kb}
-\end{alignat*}
-In messages~1 and~3, Charlie removes the encryption using his private
-key and re-encrypts Alice's messages using Bob's public key. Bob is
-left thinking he has run the protocol with Alice, which was not
-Alice's intention, and Bob is unaware that the ``secret'' nonces are
-known to Charlie. This is a typical man-in-the-middle attack launched
-by an insider.
-
-Whether this counts as an attack has been disputed. In protocols of this
-type, we normally assume that the other party is honest. To be honest
-means to obey the protocol rules, so Alice's running the protocol with
-Charlie does not make her dishonest, just careless. After Lowe's
-attack, Alice has no grounds for complaint: this protocol does not have to
-guarantee anything if you run it with a bad person. Bob does have
-grounds for complaint, however: the protocol tells him that he is
-communicating with Alice (who is honest) but it does not guarantee
-secrecy of the nonces.
-
-Lowe also suggested a correction, namely to include Bob's name in
-message~2:
-\begin{alignat*}{2}
- &1.&\quad A\to B &: \comp{Na,A}\sb{Kb} \\
- &2.&\quad B\to A &: \comp{Na,Nb,B}\sb{Ka} \\
- &3.&\quad A\to B &: \comp{Nb}\sb{Kb}
-\end{alignat*}
-If Charlie tries the same attack, Alice will receive the message
-$\comp{Na,Nb,B}\sb{Ka}$ when she was expecting to receive
-$\comp{Na,Nb,C}\sb{Ka}$. She will abandon the run, and eventually so
-will Bob. Below, we shall look at parts of this protocol's correctness
-proof.
-
-In ground-breaking work, Lowe~\cite{lowe-fdr}\index{Lowe, Gavin|)}
-showed how such attacks
-could be found automatically using a model checker. An alternative,
-which we shall examine below, is to prove protocols correct. Proofs
-can be done under more realistic assumptions because our model does
-not have to be finite. The strategy is to formalize the operational
-semantics of the system and to prove security properties using rule
-induction.%
-\index{Needham-Schroeder protocol|)}
-
-
-\input{document/Message}
-\input{document/Event}
-\input{document/Public}
-\input{document/NS_Public}
--- a/doc-src/TutorialI/Recdef/document/Induction.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,121 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Induction}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\begin{isamarkuptext}%
-Assuming we have defined our function such that Isabelle could prove
-termination and that the recursion equations (or some suitable derived
-equations) are simplification rules, we might like to prove something about
-our function. Since the function is recursive, the natural proof principle is
-again induction. But this time the structural form of induction that comes
-with datatypes is unlikely to work well --- otherwise we could have defined the
-function by \isacommand{primrec}. Therefore \isacommand{recdef} automatically
-proves a suitable induction rule $f$\isa{{\isachardot}induct} that follows the
-recursion pattern of the particular function $f$. We call this
-\textbf{recursion induction}. Roughly speaking, it
-requires you to prove for each \isacommand{recdef} equation that the property
-you are trying to establish holds for the left-hand side provided it holds
-for all recursive calls on the right-hand side. Here is a simple example
-involving the predefined \isa{map} functional on lists:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ {\isachardoublequoteopen}map\ f\ {\isacharparenleft}sep{\isacharparenleft}x{\isacharcomma}xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep{\isacharparenleft}f\ x{\isacharcomma}\ map\ f\ xs{\isacharparenright}{\isachardoublequoteclose}%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\begin{isamarkuptxt}%
-\noindent
-Note that \isa{map\ f\ xs}
-is the result of applying \isa{f} to all elements of \isa{xs}. We prove
-this lemma by recursion induction over \isa{sep}:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-{\isacharparenleft}induct{\isacharunderscore}tac\ x\ xs\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}%
-\begin{isamarkuptxt}%
-\noindent
-The resulting proof state has three subgoals corresponding to the three
-clauses for \isa{sep}:
-\begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a{\isachardot}\ map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline
-\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ x{\isachardot}\ map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\isanewline
-\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}a\ x\ y\ zs{\isachardot}\isanewline
-\isaindent{\ {\isadigit{3}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharparenleft}y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline
-\isaindent{\ {\isadigit{3}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}%
-\end{isabelle}
-The rest is pure simplification:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-\ simp{\isacharunderscore}all\isanewline
-\isacommand{done}\isamarkupfalse%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-Try proving the above lemma by structural induction, and you find that you
-need an additional case distinction. What is worse, the names of variables
-are invented by Isabelle and have nothing to do with the names in the
-definition of \isa{sep}.
-
-In general, the format of invoking recursion induction is
-\begin{quote}
-\isacommand{apply}\isa{{\isacharparenleft}induct{\isacharunderscore}tac} $x@1 \dots x@n$ \isa{rule{\isacharcolon}} $f$\isa{{\isachardot}induct{\isacharparenright}}
-\end{quote}\index{*induct_tac (method)}%
-where $x@1~\dots~x@n$ is a list of free variables in the subgoal and $f$ the
-name of a function that takes an $n$-tuple. Usually the subgoal will
-contain the term $f(x@1,\dots,x@n)$ but this need not be the case. The
-induction rules do not mention $f$ at all. Here is \isa{sep{\isachardot}induct}:
-\begin{isabelle}
-{\isasymlbrakk}~{\isasymAnd}a.~P~a~[];\isanewline
-~~{\isasymAnd}a~x.~P~a~[x];\isanewline
-~~{\isasymAnd}a~x~y~zs.~P~a~(y~\#~zs)~{\isasymLongrightarrow}~P~a~(x~\#~y~\#~zs){\isasymrbrakk}\isanewline
-{\isasymLongrightarrow}~P~u~v%
-\end{isabelle}
-It merely says that in order to prove a property \isa{P} of \isa{u} and
-\isa{v} you need to prove it for the three cases where \isa{v} is the
-empty list, the singleton list, and the list with at least two elements.
-The final case has an induction hypothesis: you may assume that \isa{P}
-holds for the tail of that list.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/Recdef/document/Nested0.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,55 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Nested{\isadigit{0}}}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\begin{isamarkuptext}%
-\index{datatypes!nested}%
-In \S\ref{sec:nested-datatype} we defined the datatype of terms%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{datatype}\isamarkupfalse%
-\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}{\isachardoublequoteopen}term{\isachardoublequoteclose}\ {\isacharequal}\ Var\ {\isacharprime}a\ {\isacharbar}\ App\ {\isacharprime}b\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term\ list{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent
-and closed with the observation that the associated schema for the definition
-of primitive recursive functions leads to overly verbose definitions. Moreover,
-if you have worked exercise~\ref{ex:trev-trev} you will have noticed that
-you needed to declare essentially the same function as \isa{rev}
-and prove many standard properties of list reversal all over again.
-We will now show you how \isacommand{recdef} can simplify
-definitions and proofs about nested recursive datatypes. As an example we
-choose exercise~\ref{ex:trev-trev}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ trev\ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term{\isachardoublequoteclose}%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/Recdef/document/Nested1.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,75 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Nested{\isadigit{1}}}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\begin{isamarkuptext}%
-\noindent
-Although the definition of \isa{trev} below is quite natural, we will have
-to overcome a minor difficulty in convincing Isabelle of its termination.
-It is precisely this difficulty that is the \textit{raison d'\^etre} of
-this subsection.
-
-Defining \isa{trev} by \isacommand{recdef} rather than \isacommand{primrec}
-simplifies matters because we are now free to use the recursion equation
-suggested at the end of \S\ref{sec:nested-datatype}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{recdef}\isamarkupfalse%
-\ trev\ {\isachardoublequoteopen}measure\ size{\isachardoublequoteclose}\isanewline
-\ {\isachardoublequoteopen}trev\ {\isacharparenleft}Var\ x{\isacharparenright}\ \ \ \ {\isacharequal}\ Var\ x{\isachardoublequoteclose}\isanewline
-\ {\isachardoublequoteopen}trev\ {\isacharparenleft}App\ f\ ts{\isacharparenright}\ {\isacharequal}\ App\ f\ {\isacharparenleft}rev{\isacharparenleft}map\ trev\ ts{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent
-Remember that function \isa{size} is defined for each \isacommand{datatype}.
-However, the definition does not succeed. Isabelle complains about an
-unproved termination condition
-\begin{isabelle}%
-\ \ \ \ \ t\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ size\ t\ {\isacharless}\ Suc\ {\isacharparenleft}term{\isacharunderscore}list{\isacharunderscore}size\ ts{\isacharparenright}%
-\end{isabelle}
-where \isa{set} returns the set of elements of a list
-and \isa{term{\isacharunderscore}list{\isacharunderscore}size\ {\isacharcolon}{\isacharcolon}\ term\ list\ {\isasymRightarrow}\ nat} is an auxiliary
-function automatically defined by Isabelle
-(while processing the declaration of \isa{term}). Why does the
-recursive call of \isa{trev} lead to this
-condition? Because \isacommand{recdef} knows that \isa{map}
-will apply \isa{trev} only to elements of \isa{ts}. Thus the
-condition expresses that the size of the argument \isa{t\ {\isasymin}\ set\ ts} of any
-recursive call of \isa{trev} is strictly less than \isa{size\ {\isacharparenleft}App\ f\ ts{\isacharparenright}},
-which equals \isa{Suc\ {\isacharparenleft}term{\isacharunderscore}list{\isacharunderscore}size\ ts{\isacharparenright}}. We will now prove the termination condition and
-continue with our definition. Below we return to the question of how
-\isacommand{recdef} knows about \isa{map}.
-
-The termination condition is easily proved by induction:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/Recdef/document/Nested2.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,152 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Nested{\isadigit{2}}}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-\isanewline
-%
-\endisadelimtheory
-\isacommand{lemma}\isamarkupfalse%
-\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}t\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ size\ t\ {\isacharless}\ Suc{\isacharparenleft}term{\isacharunderscore}list{\isacharunderscore}size\ ts{\isacharparenright}{\isachardoublequoteclose}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{by}\isamarkupfalse%
-{\isacharparenleft}induct{\isacharunderscore}tac\ ts{\isacharcomma}\ auto{\isacharparenright}%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\noindent
-By making this theorem a simplification rule, \isacommand{recdef}
-applies it automatically and the definition of \isa{trev}
-succeeds now. As a reward for our effort, we can now prove the desired
-lemma directly. We no longer need the verbose
-induction schema for type \isa{term} and can use the simpler one arising from
-\isa{trev}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ {\isachardoublequoteopen}trev{\isacharparenleft}trev\ t{\isacharparenright}\ {\isacharequal}\ t{\isachardoublequoteclose}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{apply}\isamarkupfalse%
-{\isacharparenleft}induct{\isacharunderscore}tac\ t\ rule{\isacharcolon}\ trev{\isachardot}induct{\isacharparenright}%
-\begin{isamarkuptxt}%
-\begin{isabelle}%
-\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ trev\ {\isacharparenleft}trev\ {\isacharparenleft}Var\ x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ Var\ x\isanewline
-\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}f\ ts{\isachardot}\isanewline
-\isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }{\isasymforall}x{\isachardot}\ x\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ trev\ {\isacharparenleft}trev\ x{\isacharparenright}\ {\isacharequal}\ x\ {\isasymLongrightarrow}\isanewline
-\isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }trev\ {\isacharparenleft}trev\ {\isacharparenleft}App\ f\ ts{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ App\ f\ ts%
-\end{isabelle}
-Both the base case and the induction step fall to simplification:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{by}\isamarkupfalse%
-{\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ rev{\isacharunderscore}map\ sym{\isacharbrackleft}OF\ map{\isacharunderscore}compose{\isacharbrackright}\ cong{\isacharcolon}\ map{\isacharunderscore}cong{\isacharparenright}%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\noindent
-If the proof of the induction step mystifies you, we recommend that you go through
-the chain of simplification steps in detail; you will probably need the help of
-\isa{trace{\isacharunderscore}simp}. Theorem \isa{map{\isacharunderscore}cong} is discussed below.
-%\begin{quote}
-%{term[display]"trev(trev(App f ts))"}\\
-%{term[display]"App f (rev(map trev (rev(map trev ts))))"}\\
-%{term[display]"App f (map trev (rev(rev(map trev ts))))"}\\
-%{term[display]"App f (map trev (map trev ts))"}\\
-%{term[display]"App f (map (trev o trev) ts)"}\\
-%{term[display]"App f (map (%x. x) ts)"}\\
-%{term[display]"App f ts"}
-%\end{quote}
-
-The definition of \isa{trev} above is superior to the one in
-\S\ref{sec:nested-datatype} because it uses \isa{rev}
-and lets us use existing facts such as \hbox{\isa{rev\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs}}.
-Thus this proof is a good example of an important principle:
-\begin{quote}
-\emph{Chose your definitions carefully\\
-because they determine the complexity of your proofs.}
-\end{quote}
-
-Let us now return to the question of how \isacommand{recdef} can come up with
-sensible termination conditions in the presence of higher-order functions
-like \isa{map}. For a start, if nothing were known about \isa{map}, then
-\isa{map\ trev\ ts} might apply \isa{trev} to arbitrary terms, and thus
-\isacommand{recdef} would try to prove the unprovable \isa{size\ t\ {\isacharless}\ Suc\ {\isacharparenleft}term{\isacharunderscore}list{\isacharunderscore}size\ ts{\isacharparenright}}, without any assumption about \isa{t}. Therefore
-\isacommand{recdef} has been supplied with the congruence theorem
-\isa{map{\isacharunderscore}cong}:
-\begin{isabelle}%
-\ \ \ \ \ {\isasymlbrakk}xs\ {\isacharequal}\ ys{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ set\ ys\ {\isasymLongrightarrow}\ f\ x\ {\isacharequal}\ g\ x{\isasymrbrakk}\isanewline
-\isaindent{\ \ \ \ \ }{\isasymLongrightarrow}\ map\ f\ xs\ {\isacharequal}\ map\ g\ ys%
-\end{isabelle}
-Its second premise expresses that in \isa{map\ f\ xs},
-function \isa{f} is only applied to elements of list \isa{xs}. Congruence
-rules for other higher-order functions on lists are similar. If you get
-into a situation where you need to supply \isacommand{recdef} with new
-congruence rules, you can append a hint after the end of
-the recursion equations:\cmmdx{hints}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}cong{\isacharcolon}\ map{\isacharunderscore}cong{\isacharparenright}%
-\begin{isamarkuptext}%
-\noindent
-Or you can declare them globally
-by giving them the \attrdx{recdef_cong} attribute:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{declare}\isamarkupfalse%
-\ map{\isacharunderscore}cong{\isacharbrackleft}recdef{\isacharunderscore}cong{\isacharbrackright}%
-\begin{isamarkuptext}%
-The \isa{cong} and \isa{recdef{\isacharunderscore}cong} attributes are
-intentionally kept apart because they control different activities, namely
-simplification and making recursive definitions.
-%The simplifier's congruence rules cannot be used by recdef.
-%For example the weak congruence rules for if and case would prevent
-%recdef from generating sensible termination conditions.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/Recdef/document/examples.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,133 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{examples}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\begin{isamarkuptext}%
-Here is a simple example, the \rmindex{Fibonacci function}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ fib\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ fib\ {\isachardoublequoteopen}measure{\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}fib\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}fib\ {\isacharparenleft}Suc{\isacharparenleft}Suc\ x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ fib\ x\ {\isacharplus}\ fib\ {\isacharparenleft}Suc\ x{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent
-\index{measure functions}%
-The definition of \isa{fib} is accompanied by a \textbf{measure function}
-\isa{{\isasymlambda}n{\isachardot}\ n} which maps the argument of \isa{fib} to a
-natural number. The requirement is that in each equation the measure of the
-argument on the left-hand side is strictly greater than the measure of the
-argument of each recursive call. In the case of \isa{fib} this is
-obviously true because the measure function is the identity and
-\isa{Suc\ {\isacharparenleft}Suc\ x{\isacharparenright}} is strictly greater than both \isa{x} and
-\isa{Suc\ x}.
-
-Slightly more interesting is the insertion of a fixed element
-between any two elements of a list:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ sep\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ sep\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}a{\isacharcomma}xs{\isacharparenright}{\isachardot}\ length\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}sep{\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ \ \ \ \ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}sep{\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\ \ \ \ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}sep{\isacharparenleft}a{\isacharcomma}\ x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep{\isacharparenleft}a{\isacharcomma}y{\isacharhash}zs{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent
-This time the measure is the length of the list, which decreases with the
-recursive call; the first component of the argument tuple is irrelevant.
-The details of tupled $\lambda$-abstractions \isa{{\isasymlambda}{\isacharparenleft}x\isactrlsub {\isadigit{1}}{\isacharcomma}{\isasymdots}{\isacharcomma}x\isactrlsub n{\isacharparenright}} are
-explained in \S\ref{sec:products}, but for now your intuition is all you need.
-
-Pattern matching\index{pattern matching!and \isacommand{recdef}}
-need not be exhaustive:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ last\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ last\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}xs{\isachardot}\ length\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}last\ {\isacharbrackleft}x{\isacharbrackright}\ \ \ \ \ \ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}last\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ last\ {\isacharparenleft}y{\isacharhash}zs{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-Overlapping patterns are disambiguated by taking the order of equations into
-account, just as in functional programming:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ sep{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ sep{\isadigit{1}}\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}a{\isacharcomma}xs{\isacharparenright}{\isachardot}\ length\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}sep{\isadigit{1}}{\isacharparenleft}a{\isacharcomma}\ x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep{\isadigit{1}}{\isacharparenleft}a{\isacharcomma}y{\isacharhash}zs{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}sep{\isadigit{1}}{\isacharparenleft}a{\isacharcomma}\ xs{\isacharparenright}\ \ \ \ \ {\isacharequal}\ xs{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent
-To guarantee that the second equation can only be applied if the first
-one does not match, Isabelle internally replaces the second equation
-by the two possibilities that are left: \isa{sep{\isadigit{1}}\ {\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}} and
-\isa{sep{\isadigit{1}}\ {\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}}. Thus the functions \isa{sep} and
-\isa{sep{\isadigit{1}}} are identical.
-
-\begin{warn}
- \isacommand{recdef} only takes the first argument of a (curried)
- recursive function into account. This means both the termination measure
- and pattern matching can only use that first argument. In general, you will
- therefore have to combine several arguments into a tuple. In case only one
- argument is relevant for termination, you can also rearrange the order of
- arguments as in the following definition:
-\end{warn}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ sep{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ sep{\isadigit{2}}\ {\isachardoublequoteopen}measure\ length{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}sep{\isadigit{2}}\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymlambda}a{\isachardot}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep{\isadigit{2}}\ {\isacharparenleft}y{\isacharhash}zs{\isacharparenright}\ a{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}sep{\isadigit{2}}\ xs\ \ \ \ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymlambda}a{\isachardot}\ xs{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-Because of its pattern-matching syntax, \isacommand{recdef} is also useful
-for the definition of non-recursive functions, where the termination measure
-degenerates to the empty set \isa{{\isacharbraceleft}{\isacharbraceright}}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ swap{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ swap{\isadigit{1}}{\isadigit{2}}\ {\isachardoublequoteopen}{\isacharbraceleft}{\isacharbraceright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}swap{\isadigit{1}}{\isadigit{2}}\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ y{\isacharhash}x{\isacharhash}zs{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}swap{\isadigit{1}}{\isadigit{2}}\ zs\ \ \ \ \ \ \ {\isacharequal}\ zs{\isachardoublequoteclose}\isanewline
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/Recdef/document/simplification.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,171 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{simplification}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\begin{isamarkuptext}%
-Once we have proved all the termination conditions, the \isacommand{recdef}
-recursion equations become simplification rules, just as with
-\isacommand{primrec}. In most cases this works fine, but there is a subtle
-problem that must be mentioned: simplification may not
-terminate because of automatic splitting of \isa{if}.
-\index{*if expressions!splitting of}
-Let us look at an example:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ gcd\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat{\isasymtimes}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ gcd\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}{\isachardot}n{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}gcd\ {\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ n{\isacharequal}{\isadigit{0}}\ then\ m\ else\ gcd{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent
-According to the measure function, the second argument should decrease with
-each recursive call. The resulting termination condition
-\begin{isabelle}%
-\ \ \ \ \ n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ m\ mod\ n\ {\isacharless}\ n%
-\end{isabelle}
-is proved automatically because it is already present as a lemma in
-HOL\@. Thus the recursion equation becomes a simplification
-rule. Of course the equation is nonterminating if we are allowed to unfold
-the recursive call inside the \isa{else} branch, which is why programming
-languages and our simplifier don't do that. Unfortunately the simplifier does
-something else that leads to the same problem: it splits
-each \isa{if}-expression unless its
-condition simplifies to \isa{True} or \isa{False}. For
-example, simplification reduces
-\begin{isabelle}%
-\ \ \ \ \ gcd\ {\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ k%
-\end{isabelle}
-in one step to
-\begin{isabelle}%
-\ \ \ \ \ {\isacharparenleft}if\ n\ {\isacharequal}\ {\isadigit{0}}\ then\ m\ else\ gcd\ {\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ k%
-\end{isabelle}
-where the condition cannot be reduced further, and splitting leads to
-\begin{isabelle}%
-\ \ \ \ \ {\isacharparenleft}n\ {\isacharequal}\ {\isadigit{0}}\ {\isasymlongrightarrow}\ m\ {\isacharequal}\ k{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymlongrightarrow}\ gcd\ {\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}\ {\isacharequal}\ k{\isacharparenright}%
-\end{isabelle}
-Since the recursive call \isa{gcd\ {\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}} is no longer protected by
-an \isa{if}, it is unfolded again, which leads to an infinite chain of
-simplification steps. Fortunately, this problem can be avoided in many
-different ways.
-
-The most radical solution is to disable the offending theorem
-\isa{split{\isacharunderscore}if},
-as shown in \S\ref{sec:AutoCaseSplits}. However, we do not recommend this
-approach: you will often have to invoke the rule explicitly when
-\isa{if} is involved.
-
-If possible, the definition should be given by pattern matching on the left
-rather than \isa{if} on the right. In the case of \isa{gcd} the
-following alternative definition suggests itself:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ gcd{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat{\isasymtimes}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ gcd{\isadigit{1}}\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}{\isachardot}n{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}gcd{\isadigit{1}}\ {\isacharparenleft}m{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}gcd{\isadigit{1}}\ {\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ gcd{\isadigit{1}}{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent
-The order of equations is important: it hides the side condition
-\isa{n\ {\isasymnoteq}\ {\isadigit{0}}}. Unfortunately, in general the case distinction
-may not be expressible by pattern matching.
-
-A simple alternative is to replace \isa{if} by \isa{case},
-which is also available for \isa{bool} and is not split automatically:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ gcd{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat{\isasymtimes}nat\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ gcd{\isadigit{2}}\ {\isachardoublequoteopen}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}{\isachardot}n{\isacharparenright}{\isachardoublequoteclose}\isanewline
-\ \ {\isachardoublequoteopen}gcd{\isadigit{2}}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}case\ n{\isacharequal}{\isadigit{0}}\ of\ True\ {\isasymRightarrow}\ m\ {\isacharbar}\ False\ {\isasymRightarrow}\ gcd{\isadigit{2}}{\isacharparenleft}n{\isacharcomma}m\ mod\ n{\isacharparenright}{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent
-This is probably the neatest solution next to pattern matching, and it is
-always available.
-
-A final alternative is to replace the offending simplification rules by
-derived conditional ones. For \isa{gcd} it means we have to prove
-these lemmas:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}gcd\ {\isacharparenleft}m{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ m{\isachardoublequoteclose}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{apply}\isamarkupfalse%
-{\isacharparenleft}simp{\isacharparenright}\isanewline
-\isacommand{done}\isamarkupfalse%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-\isanewline
-%
-\endisadelimproof
-\isanewline
-\isacommand{lemma}\isamarkupfalse%
-\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ gcd{\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ gcd{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{apply}\isamarkupfalse%
-{\isacharparenleft}simp{\isacharparenright}\isanewline
-\isacommand{done}\isamarkupfalse%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\noindent
-Simplification terminates for these proofs because the condition of the \isa{if} simplifies to \isa{True} or \isa{False}.
-Now we can disable the original simplification rule:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{declare}\isamarkupfalse%
-\ gcd{\isachardot}simps\ {\isacharbrackleft}simp\ del{\isacharbrackright}\isanewline
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/Recdef/document/termination.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,121 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{termination}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\begin{isamarkuptext}%
-When a function~$f$ is defined via \isacommand{recdef}, Isabelle tries to prove
-its termination with the help of the user-supplied measure. Each of the examples
-above is simple enough that Isabelle can automatically prove that the
-argument's measure decreases in each recursive call. As a result,
-$f$\isa{{\isachardot}simps} will contain the defining equations (or variants derived
-from them) as theorems. For example, look (via \isacommand{thm}) at
-\isa{sep{\isachardot}simps} and \isa{sep{\isadigit{1}}{\isachardot}simps} to see that they define
-the same function. What is more, those equations are automatically declared as
-simplification rules.
-
-Isabelle may fail to prove the termination condition for some
-recursive call. Let us try to define Quicksort:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{consts}\isamarkupfalse%
-\ qs\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ list\ {\isasymRightarrow}\ nat\ list{\isachardoublequoteclose}\isanewline
-\isacommand{recdef}\isamarkupfalse%
-\ qs\ {\isachardoublequoteopen}measure\ length{\isachardoublequoteclose}\isanewline
-\ {\isachardoublequoteopen}qs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
-\ {\isachardoublequoteopen}qs{\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ y{\isasymle}x{\isacharparenright}\ xs{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharat}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ x{\isacharless}y{\isacharparenright}\ xs{\isacharparenright}{\isachardoublequoteclose}%
-\begin{isamarkuptext}%
-\noindent where \isa{filter} is predefined and \isa{filter\ P\ xs}
-is the list of elements of \isa{xs} satisfying \isa{P}.
-This definition of \isa{qs} fails, and Isabelle prints an error message
-showing you what it was unable to prove:
-\begin{isabelle}%
-\ \ \ \ \ length\ {\isacharparenleft}filter\ {\isachardot}{\isachardot}{\isachardot}\ xs{\isacharparenright}\ {\isacharless}\ Suc\ {\isacharparenleft}length\ xs{\isacharparenright}%
-\end{isabelle}
-We can either prove this as a separate lemma, or try to figure out which
-existing lemmas may help. We opt for the second alternative. The theory of
-lists contains the simplification rule \isa{length\ {\isacharparenleft}filter\ P\ xs{\isacharparenright}\ {\isasymle}\ length\ xs},
-which is what we need, provided we turn \mbox{\isa{{\isacharless}\ Suc}}
-into
-\isa{{\isasymle}} so that the rule applies. Lemma
-\isa{less{\isacharunderscore}Suc{\isacharunderscore}eq{\isacharunderscore}le} does just that: \isa{{\isacharparenleft}m\ {\isacharless}\ Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}m\ {\isasymle}\ n{\isacharparenright}}.
-
-Now we retry the above definition but supply the lemma(s) just found (or
-proved). Because \isacommand{recdef}'s termination prover involves
-simplification, we include in our second attempt a hint: the
-\attrdx{recdef_simp} attribute says to use \isa{less{\isacharunderscore}Suc{\isacharunderscore}eq{\isacharunderscore}le} as a
-simplification rule.\cmmdx{hints}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{recdef}\isamarkupfalse%
-\ qs\ {\isachardoublequoteopen}measure\ length{\isachardoublequoteclose}\isanewline
-\ {\isachardoublequoteopen}qs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
-\ {\isachardoublequoteopen}qs{\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ y{\isasymle}x{\isacharparenright}\ xs{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharat}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ x{\isacharless}y{\isacharparenright}\ xs{\isacharparenright}{\isachardoublequoteclose}\isanewline
-{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}simp{\isacharcolon}\ less{\isacharunderscore}Suc{\isacharunderscore}eq{\isacharunderscore}le{\isacharparenright}%
-\begin{isamarkuptext}%
-\noindent
-This time everything works fine. Now \isa{qs{\isachardot}simps} contains precisely
-the stated recursion equations for \isa{qs} and they have become
-simplification rules.
-Thus we can automatically prove results such as this one:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{theorem}\isamarkupfalse%
-\ {\isachardoublequoteopen}qs{\isacharbrackleft}{\isadigit{2}}{\isacharcomma}{\isadigit{3}}{\isacharcomma}{\isadigit{0}}{\isacharbrackright}\ {\isacharequal}\ qs{\isacharbrackleft}{\isadigit{3}}{\isacharcomma}{\isadigit{0}}{\isacharcomma}{\isadigit{2}}{\isacharbrackright}{\isachardoublequoteclose}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{apply}\isamarkupfalse%
-{\isacharparenleft}simp{\isacharparenright}\isanewline
-\isacommand{done}\isamarkupfalse%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\noindent
-More exciting theorems require induction, which is discussed below.
-
-If the termination proof requires a lemma that is of general use, you can
-turn it permanently into a simplification rule, in which case the above
-\isacommand{hint} is not necessary. But in the case of
-\isa{less{\isacharunderscore}Suc{\isacharunderscore}eq{\isacharunderscore}le} this would be of dubious value.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/Rules/rules.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2641 +0,0 @@
-%!TEX root = ../tutorial.tex
-\chapter{The Rules of the Game}
-\label{chap:rules}
-
-This chapter outlines the concepts and techniques that underlie reasoning
-in Isabelle. Until now, we have proved everything using only induction and
-simplification, but any serious verification project requires more elaborate
-forms of inference. The chapter also introduces the fundamentals of
-predicate logic. The first examples in this chapter will consist of
-detailed, low-level proof steps. Later, we shall see how to automate such
-reasoning using the methods
-\isa{blast},
-\isa{auto} and others. Backward or goal-directed proof is our usual style,
-but the chapter also introduces forward reasoning, where one theorem is
-transformed to yield another.
-
-\section{Natural Deduction}
-
-\index{natural deduction|(}%
-In Isabelle, proofs are constructed using inference rules. The
-most familiar inference rule is probably \emph{modus ponens}:%
-\index{modus ponens@\emph{modus ponens}}
-\[ \infer{Q}{P\imp Q & P} \]
-This rule says that from $P\imp Q$ and $P$ we may infer~$Q$.
-
-\textbf{Natural deduction} is an attempt to formalize logic in a way
-that mirrors human reasoning patterns.
-For each logical symbol (say, $\conj$), there
-are two kinds of rules: \textbf{introduction} and \textbf{elimination} rules.
-The introduction rules allow us to infer this symbol (say, to
-infer conjunctions). The elimination rules allow us to deduce
-consequences from this symbol. Ideally each rule should mention
-one symbol only. For predicate logic this can be
-done, but when users define their own concepts they typically
-have to refer to other symbols as well. It is best not to be dogmatic.
-
-Natural deduction generally deserves its name. It is easy to use. Each
-proof step consists of identifying the outermost symbol of a formula and
-applying the corresponding rule. It creates new subgoals in
-an obvious way from parts of the chosen formula. Expanding the
-definitions of constants can blow up the goal enormously. Deriving natural
-deduction rules for such constants lets us reason in terms of their key
-properties, which might otherwise be obscured by the technicalities of its
-definition. Natural deduction rules also lend themselves to automation.
-Isabelle's
-\textbf{classical reasoner} accepts any suitable collection of natural deduction
-rules and uses them to search for proofs automatically. Isabelle is designed around
-natural deduction and many of its tools use the terminology of introduction
-and elimination rules.%
-\index{natural deduction|)}
-
-
-\section{Introduction Rules}
-
-\index{introduction rules|(}%
-An introduction rule tells us when we can infer a formula
-containing a specific logical symbol. For example, the conjunction
-introduction rule says that if we have $P$ and if we have $Q$ then
-we have $P\conj Q$. In a mathematics text, it is typically shown
-like this:
-\[ \infer{P\conj Q}{P & Q} \]
-The rule introduces the conjunction
-symbol~($\conj$) in its conclusion. In Isabelle proofs we
-mainly reason backwards. When we apply this rule, the subgoal already has
-the form of a conjunction; the proof step makes this conjunction symbol
-disappear.
-
-In Isabelle notation, the rule looks like this:
-\begin{isabelle}
-\isasymlbrakk?P;\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P\ \isasymand\ ?Q\rulenamedx{conjI}
-\end{isabelle}
-Carefully examine the syntax. The premises appear to the
-left of the arrow and the conclusion to the right. The premises (if
-more than one) are grouped using the fat brackets. The question marks
-indicate \textbf{schematic variables} (also called
-\textbf{unknowns}):\index{unknowns|bold} they may
-be replaced by arbitrary formulas. If we use the rule backwards, Isabelle
-tries to unify the current subgoal with the conclusion of the rule, which
-has the form \isa{?P\ \isasymand\ ?Q}. (Unification is discussed below,
-{\S}\ref{sec:unification}.) If successful,
-it yields new subgoals given by the formulas assigned to
-\isa{?P} and \isa{?Q}.
-
-The following trivial proof illustrates how rules work. It also introduces a
-style of indentation. If a command adds a new subgoal, then the next
-command's indentation is increased by one space; if it proves a subgoal, then
-the indentation is reduced. This provides the reader with hints about the
-subgoal structure.
-\begin{isabelle}
-\isacommand{lemma}\ conj_rule:\ "\isasymlbrakk P;\
-Q\isasymrbrakk\ \isasymLongrightarrow\ P\ \isasymand\
-(Q\ \isasymand\ P)"\isanewline
-\isacommand{apply}\ (rule\ conjI)\isanewline
-\ \isacommand{apply}\ assumption\isanewline
-\isacommand{apply}\ (rule\ conjI)\isanewline
-\ \isacommand{apply}\ assumption\isanewline
-\isacommand{apply}\ assumption
-\end{isabelle}
-At the start, Isabelle presents
-us with the assumptions (\isa{P} and~\isa{Q}) and with the goal to be proved,
-\isa{P\ \isasymand\
-(Q\ \isasymand\ P)}. We are working backwards, so when we
-apply conjunction introduction, the rule removes the outermost occurrence
-of the \isa{\isasymand} symbol. To apply a rule to a subgoal, we apply
-the proof method \isa{rule} --- here with \isa{conjI}, the conjunction
-introduction rule.
-\begin{isabelle}
-%\isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P\ \isasymand\ Q\
-%\isasymand\ P\isanewline
-\ 1.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P\isanewline
-\ 2.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ Q\ \isasymand\ P
-\end{isabelle}
-Isabelle leaves two new subgoals: the two halves of the original conjunction.
-The first is simply \isa{P}, which is trivial, since \isa{P} is among
-the assumptions. We can apply the \methdx{assumption}
-method, which proves a subgoal by finding a matching assumption.
-\begin{isabelle}
-\ 1.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\
-Q\ \isasymand\ P
-\end{isabelle}
-We are left with the subgoal of proving
-\isa{Q\ \isasymand\ P} from the assumptions \isa{P} and~\isa{Q}. We apply
-\isa{rule conjI} again.
-\begin{isabelle}
-\ 1.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ Q\isanewline
-\ 2.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P
-\end{isabelle}
-We are left with two new subgoals, \isa{Q} and~\isa{P}, each of which can be proved
-using the \isa{assumption} method.%
-\index{introduction rules|)}
-
-
-\section{Elimination Rules}
-
-\index{elimination rules|(}%
-Elimination rules work in the opposite direction from introduction
-rules. In the case of conjunction, there are two such rules.
-From $P\conj Q$ we infer $P$. also, from $P\conj Q$
-we infer $Q$:
-\[ \infer{P}{P\conj Q} \qquad \infer{Q}{P\conj Q} \]
-
-Now consider disjunction. There are two introduction rules, which resemble inverted forms of the
-conjunction elimination rules:
-\[ \infer{P\disj Q}{P} \qquad \infer{P\disj Q}{Q} \]
-
-What is the disjunction elimination rule? The situation is rather different from
-conjunction. From $P\disj Q$ we cannot conclude that $P$ is true and we
-cannot conclude that $Q$ is true; there are no direct
-elimination rules of the sort that we have seen for conjunction. Instead,
-there is an elimination rule that works indirectly. If we are trying to prove
-something else, say $R$, and we know that $P\disj Q$ holds, then we have to consider
-two cases. We can assume that $P$ is true and prove $R$ and then assume that $Q$ is
-true and prove $R$ a second time. Here we see a fundamental concept used in natural
-deduction: that of the \textbf{assumptions}. We have to prove $R$ twice, under
-different assumptions. The assumptions are local to these subproofs and are visible
-nowhere else.
-
-In a logic text, the disjunction elimination rule might be shown
-like this:
-\[ \infer{R}{P\disj Q & \infer*{R}{[P]} & \infer*{R}{[Q]}} \]
-The assumptions $[P]$ and $[Q]$ are bracketed
-to emphasize that they are local to their subproofs. In Isabelle
-notation, the already-familiar \isa{\isasymLongrightarrow} syntax serves the
-same purpose:
-\begin{isabelle}
-\isasymlbrakk?P\ \isasymor\ ?Q;\ ?P\ \isasymLongrightarrow\ ?R;\ ?Q\ \isasymLongrightarrow\ ?R\isasymrbrakk\ \isasymLongrightarrow\ ?R\rulenamedx{disjE}
-\end{isabelle}
-When we use this sort of elimination rule backwards, it produces
-a case split. (We have seen this before, in proofs by induction.) The following proof
-illustrates the use of disjunction elimination.
-\begin{isabelle}
-\isacommand{lemma}\ disj_swap:\ "P\ \isasymor\ Q\
-\isasymLongrightarrow\ Q\ \isasymor\ P"\isanewline
-\isacommand{apply}\ (erule\ disjE)\isanewline
-\ \isacommand{apply}\ (rule\ disjI2)\isanewline
-\ \isacommand{apply}\ assumption\isanewline
-\isacommand{apply}\ (rule\ disjI1)\isanewline
-\isacommand{apply}\ assumption
-\end{isabelle}
-We assume \isa{P\ \isasymor\ Q} and
-must prove \isa{Q\ \isasymor\ P}\@. Our first step uses the disjunction
-elimination rule, \isa{disjE}\@. We invoke it using \methdx{erule}, a
-method designed to work with elimination rules. It looks for an assumption that
-matches the rule's first premise. It deletes the matching assumption,
-regards the first premise as proved and returns subgoals corresponding to
-the remaining premises. When we apply \isa{erule} to \isa{disjE}, only two
-subgoals result. This is better than applying it using \isa{rule}
-to get three subgoals, then proving the first by assumption: the other
-subgoals would have the redundant assumption
-\hbox{\isa{P\ \isasymor\ Q}}.
-Most of the time, \isa{erule} is the best way to use elimination rules, since it
-replaces an assumption by its subformulas; only rarely does the original
-assumption remain useful.
-
-\begin{isabelle}
-%P\ \isasymor\ Q\ \isasymLongrightarrow\ Q\ \isasymor\ P\isanewline
-\ 1.\ P\ \isasymLongrightarrow\ Q\ \isasymor\ P\isanewline
-\ 2.\ Q\ \isasymLongrightarrow\ Q\ \isasymor\ P
-\end{isabelle}
-These are the two subgoals returned by \isa{erule}. The first assumes
-\isa{P} and the second assumes \isa{Q}. Tackling the first subgoal, we
-need to show \isa{Q\ \isasymor\ P}\@. The second introduction rule
-(\isa{disjI2}) can reduce this to \isa{P}, which matches the assumption.
-So, we apply the
-\isa{rule} method with \isa{disjI2} \ldots
-\begin{isabelle}
-\ 1.\ P\ \isasymLongrightarrow\ P\isanewline
-\ 2.\ Q\ \isasymLongrightarrow\ Q\ \isasymor\ P
-\end{isabelle}
-\ldots and finish off with the \isa{assumption}
-method. We are left with the other subgoal, which
-assumes \isa{Q}.
-\begin{isabelle}
-\ 1.\ Q\ \isasymLongrightarrow\ Q\ \isasymor\ P
-\end{isabelle}
-Its proof is similar, using the introduction
-rule \isa{disjI1}.
-
-The result of this proof is a new inference rule \isa{disj_swap}, which is neither
-an introduction nor an elimination rule, but which might
-be useful. We can use it to replace any goal of the form $Q\disj P$
-by one of the form $P\disj Q$.%
-\index{elimination rules|)}
-
-
-\section{Destruction Rules: Some Examples}
-
-\index{destruction rules|(}%
-Now let us examine the analogous proof for conjunction.
-\begin{isabelle}
-\isacommand{lemma}\ conj_swap:\ "P\ \isasymand\ Q\ \isasymLongrightarrow\ Q\ \isasymand\ P"\isanewline
-\isacommand{apply}\ (rule\ conjI)\isanewline
-\ \isacommand{apply}\ (drule\ conjunct2)\isanewline
-\ \isacommand{apply}\ assumption\isanewline
-\isacommand{apply}\ (drule\ conjunct1)\isanewline
-\isacommand{apply}\ assumption
-\end{isabelle}
-Recall that the conjunction elimination rules --- whose Isabelle names are
-\isa{conjunct1} and \isa{conjunct2} --- simply return the first or second half
-of a conjunction. Rules of this sort (where the conclusion is a subformula of a
-premise) are called \textbf{destruction} rules because they take apart and destroy
-a premise.%
-\footnote{This Isabelle terminology has no counterpart in standard logic texts,
-although the distinction between the two forms of elimination rule is well known.
-Girard \cite[page 74]{girard89},\index{Girard, Jean-Yves|fnote}
-for example, writes ``The elimination rules
-[for $\disj$ and $\exists$] are very
-bad. What is catastrophic about them is the parasitic presence of a formula [$R$]
-which has no structural link with the formula which is eliminated.''}
-
-The first proof step applies conjunction introduction, leaving
-two subgoals:
-\begin{isabelle}
-%P\ \isasymand\ Q\ \isasymLongrightarrow\ Q\ \isasymand\ P\isanewline
-\ 1.\ P\ \isasymand\ Q\ \isasymLongrightarrow\ Q\isanewline
-\ 2.\ P\ \isasymand\ Q\ \isasymLongrightarrow\ P
-\end{isabelle}
-
-To invoke the elimination rule, we apply a new method, \isa{drule}.
-Think of the \isa{d} as standing for \textbf{destruction} (or \textbf{direct}, if
-you prefer). Applying the
-second conjunction rule using \isa{drule} replaces the assumption
-\isa{P\ \isasymand\ Q} by \isa{Q}.
-\begin{isabelle}
-\ 1.\ Q\ \isasymLongrightarrow\ Q\isanewline
-\ 2.\ P\ \isasymand\ Q\ \isasymLongrightarrow\ P
-\end{isabelle}
-The resulting subgoal can be proved by applying \isa{assumption}.
-The other subgoal is similarly proved, using the \isa{conjunct1} rule and the
-\isa{assumption} method.
-
-Choosing among the methods \isa{rule}, \isa{erule} and \isa{drule} is up to
-you. Isabelle does not attempt to work out whether a rule
-is an introduction rule or an elimination rule. The
-method determines how the rule will be interpreted. Many rules
-can be used in more than one way. For example, \isa{disj_swap} can
-be applied to assumptions as well as to goals; it replaces any
-assumption of the form
-$P\disj Q$ by a one of the form $Q\disj P$.
-
-Destruction rules are simpler in form than indirect rules such as \isa{disjE},
-but they can be inconvenient. Each of the conjunction rules discards half
-of the formula, when usually we want to take both parts of the conjunction as new
-assumptions. The easiest way to do so is by using an
-alternative conjunction elimination rule that resembles \isa{disjE}\@. It is
-seldom, if ever, seen in logic books. In Isabelle syntax it looks like this:
-\begin{isabelle}
-\isasymlbrakk?P\ \isasymand\ ?Q;\ \isasymlbrakk?P;\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?R\isasymrbrakk\ \isasymLongrightarrow\ ?R\rulenamedx{conjE}
-\end{isabelle}
-\index{destruction rules|)}
-
-\begin{exercise}
-Use the rule \isa{conjE} to shorten the proof above.
-\end{exercise}
-
-
-\section{Implication}
-
-\index{implication|(}%
-At the start of this chapter, we saw the rule \emph{modus ponens}. It is, in fact,
-a destruction rule. The matching introduction rule looks like this
-in Isabelle:
-\begin{isabelle}
-(?P\ \isasymLongrightarrow\ ?Q)\ \isasymLongrightarrow\ ?P\
-\isasymlongrightarrow\ ?Q\rulenamedx{impI}
-\end{isabelle}
-And this is \emph{modus ponens}\index{modus ponens@\emph{modus ponens}}:
-\begin{isabelle}
-\isasymlbrakk?P\ \isasymlongrightarrow\ ?Q;\ ?P\isasymrbrakk\
-\isasymLongrightarrow\ ?Q
-\rulenamedx{mp}
-\end{isabelle}
-
-Here is a proof using the implication rules. This
-lemma performs a sort of uncurrying, replacing the two antecedents
-of a nested implication by a conjunction. The proof illustrates
-how assumptions work. At each proof step, the subgoals inherit the previous
-assumptions, perhaps with additions or deletions. Rules such as
-\isa{impI} and \isa{disjE} add assumptions, while applying \isa{erule} or
-\isa{drule} deletes the matching assumption.
-\begin{isabelle}
-\isacommand{lemma}\ imp_uncurry:\
-"P\ \isasymlongrightarrow\ (Q\
-\isasymlongrightarrow\ R)\ \isasymLongrightarrow\ P\
-\isasymand\ Q\ \isasymlongrightarrow\
-R"\isanewline
-\isacommand{apply}\ (rule\ impI)\isanewline
-\isacommand{apply}\ (erule\ conjE)\isanewline
-\isacommand{apply}\ (drule\ mp)\isanewline
-\ \isacommand{apply}\ assumption\isanewline
-\isacommand{apply}\ (drule\ mp)\isanewline
-\ \ \isacommand{apply}\ assumption\isanewline
-\ \isacommand{apply}\ assumption
-\end{isabelle}
-First, we state the lemma and apply implication introduction (\isa{rule impI}),
-which moves the conjunction to the assumptions.
-\begin{isabelle}
-%P\ \isasymlongrightarrow\ Q\ \isasymlongrightarrow\ R\ \isasymLongrightarrow\ P\
-%\isasymand\ Q\ \isasymlongrightarrow\ R\isanewline
-\ 1.\ \isasymlbrakk P\ \isasymlongrightarrow\ Q\ \isasymlongrightarrow\ R;\ P\ \isasymand\ Q\isasymrbrakk\ \isasymLongrightarrow\ R
-\end{isabelle}
-Next, we apply conjunction elimination (\isa{erule conjE}), which splits this
-conjunction into two parts.
-\begin{isabelle}
-\ 1.\ \isasymlbrakk P\ \isasymlongrightarrow\ Q\ \isasymlongrightarrow\ R;\ P;\
-Q\isasymrbrakk\ \isasymLongrightarrow\ R
-\end{isabelle}
-Now, we work on the assumption \isa{P\ \isasymlongrightarrow\ (Q\
-\isasymlongrightarrow\ R)}, where the parentheses have been inserted for
-clarity. The nested implication requires two applications of
-\textit{modus ponens}: \isa{drule mp}. The first use yields the
-implication \isa{Q\
-\isasymlongrightarrow\ R}, but first we must prove the extra subgoal
-\isa{P}, which we do by assumption.
-\begin{isabelle}
-\ 1.\ \isasymlbrakk P;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P\isanewline
-\ 2.\ \isasymlbrakk P;\ Q;\ Q\ \isasymlongrightarrow\ R\isasymrbrakk\ \isasymLongrightarrow\ R
-\end{isabelle}
-Repeating these steps for \isa{Q\
-\isasymlongrightarrow\ R} yields the conclusion we seek, namely~\isa{R}.
-\begin{isabelle}
-\ 1.\ \isasymlbrakk P;\ Q;\ Q\ \isasymlongrightarrow\ R\isasymrbrakk\
-\isasymLongrightarrow\ R
-\end{isabelle}
-
-The symbols \isa{\isasymLongrightarrow} and \isa{\isasymlongrightarrow}
-both stand for implication, but they differ in many respects. Isabelle
-uses \isa{\isasymLongrightarrow} to express inference rules; the symbol is
-built-in and Isabelle's inference mechanisms treat it specially. On the
-other hand, \isa{\isasymlongrightarrow} is just one of the many connectives
-available in higher-order logic. We reason about it using inference rules
-such as \isa{impI} and \isa{mp}, just as we reason about the other
-connectives. You will have to use \isa{\isasymlongrightarrow} in any
-context that requires a formula of higher-order logic. Use
-\isa{\isasymLongrightarrow} to separate a theorem's preconditions from its
-conclusion.%
-\index{implication|)}
-
-\medskip
-\index{by@\isacommand{by} (command)|(}%
-The \isacommand{by} command is useful for proofs like these that use
-\isa{assumption} heavily. It executes an
-\isacommand{apply} command, then tries to prove all remaining subgoals using
-\isa{assumption}. Since (if successful) it ends the proof, it also replaces the
-\isacommand{done} symbol. For example, the proof above can be shortened:
-\begin{isabelle}
-\isacommand{lemma}\ imp_uncurry:\
-"P\ \isasymlongrightarrow\ (Q\
-\isasymlongrightarrow\ R)\ \isasymLongrightarrow\ P\
-\isasymand\ Q\ \isasymlongrightarrow\
-R"\isanewline
-\isacommand{apply}\ (rule\ impI)\isanewline
-\isacommand{apply}\ (erule\ conjE)\isanewline
-\isacommand{apply}\ (drule\ mp)\isanewline
-\ \isacommand{apply}\ assumption\isanewline
-\isacommand{by}\ (drule\ mp)
-\end{isabelle}
-We could use \isacommand{by} to replace the final \isacommand{apply} and
-\isacommand{done} in any proof, but typically we use it
-to eliminate calls to \isa{assumption}. It is also a nice way of expressing a
-one-line proof.%
-\index{by@\isacommand{by} (command)|)}
-
-
-
-\section{Negation}
-
-\index{negation|(}%
-Negation causes surprising complexity in proofs. Its natural
-deduction rules are straightforward, but additional rules seem
-necessary in order to handle negated assumptions gracefully. This section
-also illustrates the \isa{intro} method: a convenient way of
-applying introduction rules.
-
-Negation introduction deduces $\lnot P$ if assuming $P$ leads to a
-contradiction. Negation elimination deduces any formula in the
-presence of $\lnot P$ together with~$P$:
-\begin{isabelle}
-(?P\ \isasymLongrightarrow\ False)\ \isasymLongrightarrow\ \isasymnot\ ?P%
-\rulenamedx{notI}\isanewline
-\isasymlbrakk{\isasymnot}\ ?P;\ ?P\isasymrbrakk\ \isasymLongrightarrow\ ?R%
-\rulenamedx{notE}
-\end{isabelle}
-%
-Classical logic allows us to assume $\lnot P$
-when attempting to prove~$P$:
-\begin{isabelle}
-(\isasymnot\ ?P\ \isasymLongrightarrow\ ?P)\ \isasymLongrightarrow\ ?P%
-\rulenamedx{classical}
-\end{isabelle}
-
-\index{contrapositives|(}%
-The implications $P\imp Q$ and $\lnot Q\imp\lnot P$ are logically
-equivalent, and each is called the
-\textbf{contrapositive} of the other. Four further rules support
-reasoning about contrapositives. They differ in the placement of the
-negation symbols:
-\begin{isabelle}
-\isasymlbrakk?Q;\ \isasymnot\ ?P\ \isasymLongrightarrow\ \isasymnot\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P%
-\rulename{contrapos_pp}\isanewline
-\isasymlbrakk?Q;\ ?P\ \isasymLongrightarrow\ \isasymnot\ ?Q\isasymrbrakk\ \isasymLongrightarrow\
-\isasymnot\ ?P%
-\rulename{contrapos_pn}\isanewline
-\isasymlbrakk{\isasymnot}\ ?Q;\ \isasymnot\ ?P\ \isasymLongrightarrow\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P%
-\rulename{contrapos_np}\isanewline
-\isasymlbrakk{\isasymnot}\ ?Q;\ ?P\ \isasymLongrightarrow\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ \isasymnot\ ?P%
-\rulename{contrapos_nn}
-\end{isabelle}
-%
-These rules are typically applied using the \isa{erule} method, where
-their effect is to form a contrapositive from an
-assumption and the goal's conclusion.%
-\index{contrapositives|)}
-
-The most important of these is \isa{contrapos_np}. It is useful
-for applying introduction rules to negated assumptions. For instance,
-the assumption $\lnot(P\imp Q)$ is equivalent to the conclusion $P\imp Q$ and we
-might want to use conjunction introduction on it.
-Before we can do so, we must move that assumption so that it
-becomes the conclusion. The following proof demonstrates this
-technique:
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk{\isasymnot}(P{\isasymlongrightarrow}Q);\
-\isasymnot(R{\isasymlongrightarrow}Q)\isasymrbrakk\ \isasymLongrightarrow\
-R"\isanewline
-\isacommand{apply}\ (erule_tac\ Q = "R{\isasymlongrightarrow}Q"\ \isakeyword{in}\
-contrapos_np)\isanewline
-\isacommand{apply}\ (intro\ impI)\isanewline
-\isacommand{by}\ (erule\ notE)
-\end{isabelle}
-%
-There are two negated assumptions and we need to exchange the conclusion with the
-second one. The method \isa{erule contrapos_np} would select the first assumption,
-which we do not want. So we specify the desired assumption explicitly
-using a new method, \isa{erule_tac}. This is the resulting subgoal:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk{\isasymnot}\ (P\ \isasymlongrightarrow\ Q);\ \isasymnot\
-R\isasymrbrakk\ \isasymLongrightarrow\ R\ \isasymlongrightarrow\ Q%
-\end{isabelle}
-The former conclusion, namely \isa{R}, now appears negated among the assumptions,
-while the negated formula \isa{R\ \isasymlongrightarrow\ Q} becomes the new
-conclusion.
-
-We can now apply introduction rules. We use the \methdx{intro} method, which
-repeatedly applies the given introduction rules. Here its effect is equivalent
-to \isa{rule impI}.
-\begin{isabelle}
-\ 1.\ \isasymlbrakk{\isasymnot}\ (P\ \isasymlongrightarrow\ Q);\ \isasymnot\ R;\
-R\isasymrbrakk\ \isasymLongrightarrow\ Q%
-\end{isabelle}
-We can see a contradiction in the form of assumptions \isa{\isasymnot\ R}
-and~\isa{R}, which suggests using negation elimination. If applied on its own,
-\isa{notE} will select the first negated assumption, which is useless.
-Instead, we invoke the rule using the
-\isa{by} command.
-Now when Isabelle selects the first assumption, it tries to prove \isa{P\
-\isasymlongrightarrow\ Q} and fails; it then backtracks, finds the
-assumption \isa{\isasymnot~R} and finally proves \isa{R} by assumption. That
-concludes the proof.
-
-\medskip
-
-The following example may be skipped on a first reading. It involves a
-peculiar but important rule, a form of disjunction introduction:
-\begin{isabelle}
-(\isasymnot \ ?Q\ \isasymLongrightarrow \ ?P)\ \isasymLongrightarrow \ ?P\ \isasymor \ ?Q%
-\rulenamedx{disjCI}
-\end{isabelle}
-This rule combines the effects of \isa{disjI1} and \isa{disjI2}. Its great
-advantage is that we can remove the disjunction symbol without deciding
-which disjunction to prove. This treatment of disjunction is standard in sequent
-and tableau calculi.
-
-\begin{isabelle}
-\isacommand{lemma}\ "(P\ \isasymor\ Q)\ \isasymand\ R\
-\isasymLongrightarrow\ P\ \isasymor\ (Q\ \isasymand\ R)"\isanewline
-\isacommand{apply}\ (rule\ disjCI)\isanewline
-\isacommand{apply}\ (elim\ conjE\ disjE)\isanewline
-\ \isacommand{apply}\ assumption
-\isanewline
-\isacommand{by}\ (erule\ contrapos_np,\ rule\ conjI)
-\end{isabelle}
-%
-The first proof step to applies the introduction rules \isa{disjCI}.
-The resulting subgoal has the negative assumption
-\hbox{\isa{\isasymnot(Q\ \isasymand\ R)}}.
-
-\begin{isabelle}
-\ 1.\ \isasymlbrakk(P\ \isasymor\ Q)\ \isasymand\ R;\ \isasymnot\ (Q\ \isasymand\
-R)\isasymrbrakk\ \isasymLongrightarrow\ P%
-\end{isabelle}
-Next we apply the \isa{elim} method, which repeatedly applies
-elimination rules; here, the elimination rules given
-in the command. One of the subgoals is trivial (\isa{\isacommand{apply} assumption}),
-leaving us with one other:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk{\isasymnot}\ (Q\ \isasymand\ R);\ R;\ Q\isasymrbrakk\ \isasymLongrightarrow\ P%
-\end{isabelle}
-%
-Now we must move the formula \isa{Q\ \isasymand\ R} to be the conclusion. The
-combination
-\begin{isabelle}
-\ \ \ \ \ (erule\ contrapos_np,\ rule\ conjI)
-\end{isabelle}
-is robust: the \isa{conjI} forces the \isa{erule} to select a
-conjunction. The two subgoals are the ones we would expect from applying
-conjunction introduction to
-\isa{Q~\isasymand~R}:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk R;\ Q;\ \isasymnot\ P\isasymrbrakk\ \isasymLongrightarrow\
-Q\isanewline
-\ 2.\ \isasymlbrakk R;\ Q;\ \isasymnot\ P\isasymrbrakk\ \isasymLongrightarrow\ R%
-\end{isabelle}
-They are proved by assumption, which is implicit in the \isacommand{by}
-command.%
-\index{negation|)}
-
-
-\section{Interlude: the Basic Methods for Rules}
-
-We have seen examples of many tactics that operate on individual rules. It
-may be helpful to review how they work given an arbitrary rule such as this:
-\[ \infer{Q}{P@1 & \ldots & P@n} \]
-Below, we refer to $P@1$ as the \bfindex{major premise}. This concept
-applies only to elimination and destruction rules. These rules act upon an
-instance of their major premise, typically to replace it by subformulas of itself.
-
-Suppose that the rule above is called~\isa{R}\@. Here are the basic rule
-methods, most of which we have already seen:
-\begin{itemize}
-\item
-Method \isa{rule\ R} unifies~$Q$ with the current subgoal, replacing it
-by $n$ new subgoals: instances of $P@1$, \ldots,~$P@n$.
-This is backward reasoning and is appropriate for introduction rules.
-\item
-Method \isa{erule\ R} unifies~$Q$ with the current subgoal and
-simultaneously unifies $P@1$ with some assumption. The subgoal is
-replaced by the $n-1$ new subgoals of proving
-instances of $P@2$,
-\ldots,~$P@n$, with the matching assumption deleted. It is appropriate for
-elimination rules. The method
-\isa{(rule\ R,\ assumption)} is similar, but it does not delete an
-assumption.
-\item
-Method \isa{drule\ R} unifies $P@1$ with some assumption, which it
-then deletes. The subgoal is
-replaced by the $n-1$ new subgoals of proving $P@2$, \ldots,~$P@n$; an
-$n$th subgoal is like the original one but has an additional assumption: an
-instance of~$Q$. It is appropriate for destruction rules.
-\item
-Method \isa{frule\ R} is like \isa{drule\ R} except that the matching
-assumption is not deleted. (See {\S}\ref{sec:frule} below.)
-\end{itemize}
-
-Other methods apply a rule while constraining some of its
-variables. The typical form is
-\begin{isabelle}
-\ \ \ \ \ \methdx{rule_tac}\ $v@1$ = $t@1$ \isakeyword{and} \ldots \isakeyword{and}
-$v@k$ =
-$t@k$ \isakeyword{in} R
-\end{isabelle}
-This method behaves like \isa{rule R}, while instantiating the variables
-$v@1$, \ldots,
-$v@k$ as specified. We similarly have \methdx{erule_tac}, \methdx{drule_tac} and
-\methdx{frule_tac}. These methods also let us specify which subgoal to
-operate on. By default it is the first subgoal, as with nearly all
-methods, but we can specify that rule \isa{R} should be applied to subgoal
-number~$i$:
-\begin{isabelle}
-\ \ \ \ \ rule_tac\ [$i$] R
-\end{isabelle}
-
-
-
-\section{Unification and Substitution}\label{sec:unification}
-
-\index{unification|(}%
-As we have seen, Isabelle rules involve schematic variables, which begin with
-a question mark and act as
-placeholders for terms. \textbf{Unification} --- well known to Prolog programmers --- is the act of
-making two terms identical, possibly replacing their schematic variables by
-terms. The simplest case is when the two terms are already the same. Next
-simplest is \textbf{pattern-matching}, which replaces variables in only one of the
-terms. The
-\isa{rule} method typically matches the rule's conclusion
-against the current subgoal. The
-\isa{assumption} method matches the current subgoal's conclusion
-against each of its assumptions. Unification can instantiate variables in both terms; the \isa{rule} method can do this if the goal
-itself contains schematic variables. Other occurrences of the variables in
-the rule or proof state are updated at the same time.
-
-Schematic variables in goals represent unknown terms. Given a goal such
-as $\exists x.\,P$, they let us proceed with a proof. They can be
-filled in later, sometimes in stages and often automatically.
-
-\begin{pgnote}
-If unification fails when you think it should succeed, try setting the Proof General flag \pgmenu{Isabelle} $>$ \pgmenu{Settings} $>$
-\pgmenu{Trace Unification},
-which makes Isabelle show the cause of unification failures (in Proof
-General's \pgmenu{Trace} buffer).
-\end{pgnote}
-\noindent
-For example, suppose we are trying to prove this subgoal by assumption:
-\begin{isabelle}
-\ 1.\ P\ (a,\ f\ (b,\ g\ (e,\ a),\ b),\ a)\ \isasymLongrightarrow \ P\ (a,\ f\ (b,\ g\ (c,\ a),\ b),\ a)
-\end{isabelle}
-The \isa{assumption} method having failed, we try again with the flag set:
-\begin{isabelle}
-\isacommand{apply} assumption
-\end{isabelle}
-In this trivial case, the output clearly shows that \isa{e} clashes with \isa{c}:
-\begin{isabelle}
-Clash: e =/= c
-\end{isabelle}
-
-Isabelle uses
-\textbf{higher-order} unification, which works in the
-typed $\lambda$-calculus. The procedure requires search and is potentially
-undecidable. For our purposes, however, the differences from ordinary
-unification are straightforward. It handles bound variables
-correctly, avoiding capture. The two terms
-\isa{{\isasymlambda}x.\ f(x,z)} and \isa{{\isasymlambda}y.\ f(y,z)} are
-trivially unifiable because they differ only by a bound variable renaming. The two terms \isa{{\isasymlambda}x.\ ?P} and
-\isa{{\isasymlambda}x.\ t x} are not unifiable; replacing \isa{?P} by
-\isa{t x} is forbidden because the free occurrence of~\isa{x} would become
-bound. Unfortunately, even if \isa{trace_unify_fail} is set, Isabelle displays no information about this type of failure.
-
-\begin{warn}
-Higher-order unification sometimes must invent
-$\lambda$-terms to replace function variables,
-which can lead to a combinatorial explosion. However, Isabelle proofs tend
-to involve easy cases where there are few possibilities for the
-$\lambda$-term being constructed. In the easiest case, the
-function variable is applied only to bound variables,
-as when we try to unify \isa{{\isasymlambda}x\ y.\ f(?h x y)} and
-\isa{{\isasymlambda}x\ y.\ f(x+y+a)}. The only solution is to replace
-\isa{?h} by \isa{{\isasymlambda}x\ y.\ x+y+a}. Such cases admit at most
-one unifier, like ordinary unification. A harder case is
-unifying \isa{?h a} with~\isa{a+b}; it admits two solutions for \isa{?h},
-namely \isa{{\isasymlambda}x.~a+b} and \isa{{\isasymlambda}x.~x+b}.
-Unifying \isa{?h a} with~\isa{a+a+b} admits four solutions; their number is
-exponential in the number of occurrences of~\isa{a} in the second term.
-\end{warn}
-
-
-
-\subsection{Substitution and the {\tt\slshape subst} Method}
-\label{sec:subst}
-
-\index{substitution|(}%
-Isabelle also uses function variables to express \textbf{substitution}.
-A typical substitution rule allows us to replace one term by
-another if we know that two terms are equal.
-\[ \infer{P[t/x]}{s=t & P[s/x]} \]
-The rule uses a notation for substitution: $P[t/x]$ is the result of
-replacing $x$ by~$t$ in~$P$. The rule only substitutes in the positions
-designated by~$x$. For example, it can
-derive symmetry of equality from reflexivity. Using $x=s$ for~$P$
-replaces just the first $s$ in $s=s$ by~$t$:
-\[ \infer{t=s}{s=t & \infer{s=s}{}} \]
-
-The Isabelle version of the substitution rule looks like this:
-\begin{isabelle}
-\isasymlbrakk?t\ =\ ?s;\ ?P\ ?s\isasymrbrakk\ \isasymLongrightarrow\ ?P\
-?t
-\rulenamedx{ssubst}
-\end{isabelle}
-Crucially, \isa{?P} is a function
-variable. It can be replaced by a $\lambda$-term
-with one bound variable, whose occurrences identify the places
-in which $s$ will be replaced by~$t$. The proof above requires \isa{?P}
-to be replaced by \isa{{\isasymlambda}x.~x=s}; the second premise will then
-be \isa{s=s} and the conclusion will be \isa{t=s}.
-
-The \isa{simp} method also replaces equals by equals, but the substitution
-rule gives us more control. Consider this proof:
-\begin{isabelle}
-\isacommand{lemma}\
-"\isasymlbrakk x\ =\ f\ x;\ odd(f\ x)\isasymrbrakk\ \isasymLongrightarrow\
-odd\ x"\isanewline
-\isacommand{by}\ (erule\ ssubst)
-\end{isabelle}
-%
-The assumption \isa{x\ =\ f\ x}, if used for rewriting, would loop,
-replacing \isa{x} by \isa{f x} and then by
-\isa{f(f x)} and so forth. (Here \isa{simp}
-would see the danger and would re-orient the equality, but in more complicated
-cases it can be fooled.) When we apply the substitution rule,
-Isabelle replaces every
-\isa{x} in the subgoal by \isa{f x} just once. It cannot loop. The
-resulting subgoal is trivial by assumption, so the \isacommand{by} command
-proves it implicitly.
-
-We are using the \isa{erule} method in a novel way. Hitherto,
-the conclusion of the rule was just a variable such as~\isa{?R}, but it may
-be any term. The conclusion is unified with the subgoal just as
-it would be with the \isa{rule} method. At the same time \isa{erule} looks
-for an assumption that matches the rule's first premise, as usual. With
-\isa{ssubst} the effect is to find, use and delete an equality
-assumption.
-
-The \methdx{subst} method performs individual substitutions. In simple cases,
-it closely resembles a use of the substitution rule. Suppose a
-proof has reached this point:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk P\ x\ y\ z;\ Suc\ x\ <\ y\isasymrbrakk \ \isasymLongrightarrow \ f\ z\ =\ x\ *\ y%
-\end{isabelle}
-Now we wish to apply a commutative law:
-\begin{isabelle}
-?m\ *\ ?n\ =\ ?n\ *\ ?m%
-\rulename{mult_commute}
-\end{isabelle}
-Isabelle rejects our first attempt:
-\begin{isabelle}
-apply (simp add: mult_commute)
-\end{isabelle}
-The simplifier notices the danger of looping and refuses to apply the
-rule.%
-\footnote{More precisely, it only applies such a rule if the new term is
-smaller under a specified ordering; here, \isa{x\ *\ y}
-is already smaller than
-\isa{y\ *\ x}.}
-%
-The \isa{subst} method applies \isa{mult_commute} exactly once.
-\begin{isabelle}
-\isacommand{apply}\ (subst\ mult_commute)\isanewline
-\ 1.\ \isasymlbrakk P\ x\ y\ z;\ Suc\ x\ <\ y\isasymrbrakk \
-\isasymLongrightarrow \ f\ z\ =\ y\ *\ x%
-\end{isabelle}
-As we wanted, \isa{x\ *\ y} has become \isa{y\ *\ x}.
-
-\medskip
-This use of the \methdx{subst} method has the same effect as the command
-\begin{isabelle}
-\isacommand{apply}\ (rule\ mult_commute [THEN ssubst])
-\end{isabelle}
-The attribute \isa{THEN}, which combines two rules, is described in
-{\S}\ref{sec:THEN} below. The \methdx{subst} method is more powerful than
-applying the substitution rule. It can perform substitutions in a subgoal's
-assumptions. Moreover, if the subgoal contains more than one occurrence of
-the left-hand side of the equality, the \methdx{subst} method lets us specify which occurrence should be replaced.
-
-
-\subsection{Unification and Its Pitfalls}
-
-Higher-order unification can be tricky. Here is an example, which you may
-want to skip on your first reading:
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk x\ =\
-f\ x;\ triple\ (f\ x)\ (f\ x)\ x\isasymrbrakk\
-\isasymLongrightarrow\ triple\ x\ x\ x"\isanewline
-\isacommand{apply}\ (erule\ ssubst)\isanewline
-\isacommand{back}\isanewline
-\isacommand{back}\isanewline
-\isacommand{back}\isanewline
-\isacommand{back}\isanewline
-\isacommand{apply}\ assumption\isanewline
-\isacommand{done}
-\end{isabelle}
-%
-By default, Isabelle tries to substitute for all the
-occurrences. Applying \isa{erule\ ssubst} yields this subgoal:
-\begin{isabelle}
-\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ (f\ x)\ (f\ x)\ (f\ x)
-\end{isabelle}
-The substitution should have been done in the first two occurrences
-of~\isa{x} only. Isabelle has gone too far. The \commdx{back}
-command allows us to reject this possibility and demand a new one:
-\begin{isabelle}
-\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ x\ (f\ x)\ (f\ x)
-\end{isabelle}
-%
-Now Isabelle has left the first occurrence of~\isa{x} alone. That is
-promising but it is not the desired combination. So we use \isacommand{back}
-again:
-\begin{isabelle}
-\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ (f\ x)\ x\ (f\ x)
-\end{isabelle}
-%
-This also is wrong, so we use \isacommand{back} again:
-\begin{isabelle}
-\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ x\ x\ (f\ x)
-\end{isabelle}
-%
-And this one is wrong too. Looking carefully at the series
-of alternatives, we see a binary countdown with reversed bits: 111,
-011, 101, 001. Invoke \isacommand{back} again:
-\begin{isabelle}
-\ 1.\ triple\ (f\ x)\ (f\ x)\ x\ \isasymLongrightarrow\ triple\ (f\ x)\ (f\ x)\ x%
-\end{isabelle}
-At last, we have the right combination! This goal follows by assumption.%
-\index{unification|)}
-
-\medskip
-This example shows that unification can do strange things with
-function variables. We were forced to select the right unifier using the
-\isacommand{back} command. That is all right during exploration, but \isacommand{back}
-should never appear in the final version of a proof. You can eliminate the
-need for \isacommand{back} by giving Isabelle less freedom when you apply a rule.
-
-One way to constrain the inference is by joining two methods in a
-\isacommand{apply} command. Isabelle applies the first method and then the
-second. If the second method fails then Isabelle automatically backtracks.
-This process continues until the first method produces an output that the
-second method can use. We get a one-line proof of our example:
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk x\ =\ f\ x;\ triple\ (f\ x)\ (f\ x)\ x\isasymrbrakk\
-\isasymLongrightarrow\ triple\ x\ x\ x"\isanewline
-\isacommand{apply}\ (erule\ ssubst,\ assumption)\isanewline
-\isacommand{done}
-\end{isabelle}
-
-\noindent
-The \isacommand{by} command works too, since it backtracks when
-proving subgoals by assumption:
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk x\ =\ f\ x;\ triple\ (f\ x)\ (f\ x)\ x\isasymrbrakk\
-\isasymLongrightarrow\ triple\ x\ x\ x"\isanewline
-\isacommand{by}\ (erule\ ssubst)
-\end{isabelle}
-
-
-The most general way to constrain unification is
-by instantiating variables in the rule. The method \isa{rule_tac} is
-similar to \isa{rule}, but it
-makes some of the rule's variables denote specified terms.
-Also available are {\isa{drule_tac}} and \isa{erule_tac}. Here we need
-\isa{erule_tac} since above we used \isa{erule}.
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk x\ =\ f\ x;\ triple\ (f\ x)\ (f\ x)\ x\isasymrbrakk\ \isasymLongrightarrow\ triple\ x\ x\ x"\isanewline
-\isacommand{by}\ (erule_tac\ P = "\isasymlambda u.\ triple\ u\ u\ x"\
-\isakeyword{in}\ ssubst)
-\end{isabelle}
-%
-To specify a desired substitution
-requires instantiating the variable \isa{?P} with a $\lambda$-expression.
-The bound variable occurrences in \isa{{\isasymlambda}u.\ P\ u\
-u\ x} indicate that the first two arguments have to be substituted, leaving
-the third unchanged. With this instantiation, backtracking is neither necessary
-nor possible.
-
-An alternative to \isa{rule_tac} is to use \isa{rule} with a theorem
-modified using~\isa{of}, described in
-{\S}\ref{sec:forward} below. But \isa{rule_tac}, unlike \isa{of}, can
-express instantiations that refer to
-\isasymAnd-bound variables in the current subgoal.%
-\index{substitution|)}
-
-
-\section{Quantifiers}
-
-\index{quantifiers!universal|(}%
-Quantifiers require formalizing syntactic substitution and the notion of
-arbitrary value. Consider the universal quantifier. In a logic
-book, its introduction rule looks like this:
-\[ \infer{\forall x.\,P}{P} \]
-Typically, a proviso written in English says that $x$ must not
-occur in the assumptions. This proviso guarantees that $x$ can be regarded as
-arbitrary, since it has not been assumed to satisfy any special conditions.
-Isabelle's underlying formalism, called the
-\bfindex{meta-logic}, eliminates the need for English. It provides its own
-universal quantifier (\isasymAnd) to express the notion of an arbitrary value.
-We have already seen another operator of the meta-logic, namely
-\isa\isasymLongrightarrow, which expresses inference rules and the treatment
-of assumptions. The only other operator in the meta-logic is \isa\isasymequiv,
-which can be used to define constants.
-
-\subsection{The Universal Introduction Rule}
-
-Returning to the universal quantifier, we find that having a similar quantifier
-as part of the meta-logic makes the introduction rule trivial to express:
-\begin{isabelle}
-(\isasymAnd x.\ ?P\ x)\ \isasymLongrightarrow\ {\isasymforall}x.\ ?P\ x\rulenamedx{allI}
-\end{isabelle}
-
-
-The following trivial proof demonstrates how the universal introduction
-rule works.
-\begin{isabelle}
-\isacommand{lemma}\ "{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ x"\isanewline
-\isacommand{apply}\ (rule\ allI)\isanewline
-\isacommand{by}\ (rule\ impI)
-\end{isabelle}
-The first step invokes the rule by applying the method \isa{rule allI}.
-\begin{isabelle}
-\ 1.\ \isasymAnd x.\ P\ x\ \isasymlongrightarrow\ P\ x
-\end{isabelle}
-Note that the resulting proof state has a bound variable,
-namely~\isa{x}. The rule has replaced the universal quantifier of
-higher-order logic by Isabelle's meta-level quantifier. Our goal is to
-prove
-\isa{P\ x\ \isasymlongrightarrow\ P\ x} for arbitrary~\isa{x}; it is
-an implication, so we apply the corresponding introduction rule (\isa{impI}).
-\begin{isabelle}
-\ 1.\ \isasymAnd x.\ P\ x\ \isasymLongrightarrow\ P\ x
-\end{isabelle}
-This last subgoal is implicitly proved by assumption.
-
-\subsection{The Universal Elimination Rule}
-
-Now consider universal elimination. In a logic text,
-the rule looks like this:
-\[ \infer{P[t/x]}{\forall x.\,P} \]
-The conclusion is $P$ with $t$ substituted for the variable~$x$.
-Isabelle expresses substitution using a function variable:
-\begin{isabelle}
-{\isasymforall}x.\ ?P\ x\ \isasymLongrightarrow\ ?P\ ?x\rulenamedx{spec}
-\end{isabelle}
-This destruction rule takes a
-universally quantified formula and removes the quantifier, replacing
-the bound variable \isa{x} by the schematic variable \isa{?x}. Recall that a
-schematic variable starts with a question mark and acts as a
-placeholder: it can be replaced by any term.
-
-The universal elimination rule is also
-available in the standard elimination format. Like \isa{conjE}, it never
-appears in logic books:
-\begin{isabelle}
-\isasymlbrakk \isasymforall x.\ ?P\ x;\ ?P\ ?x\ \isasymLongrightarrow \ ?R\isasymrbrakk \ \isasymLongrightarrow \ ?R%
-\rulenamedx{allE}
-\end{isabelle}
-The methods \isa{drule~spec} and \isa{erule~allE} do precisely the
-same inference.
-
-To see how $\forall$-elimination works, let us derive a rule about reducing
-the scope of a universal quantifier. In mathematical notation we write
-\[ \infer{P\imp\forall x.\,Q}{\forall x.\,P\imp Q} \]
-with the proviso ``$x$ not free in~$P$.'' Isabelle's treatment of
-substitution makes the proviso unnecessary. The conclusion is expressed as
-\isa{P\
-\isasymlongrightarrow\ ({\isasymforall}x.\ Q\ x)}. No substitution for the
-variable \isa{P} can introduce a dependence upon~\isa{x}: that would be a
-bound variable capture. Let us walk through the proof.
-\begin{isabelle}
-\isacommand{lemma}\ "(\isasymforall x.\ P\ \isasymlongrightarrow \ Q\ x)\
-\isasymLongrightarrow \ P\ \isasymlongrightarrow \ (\isasymforall x.\ Q\
-x)"
-\end{isabelle}
-First we apply implies introduction (\isa{impI}),
-which moves the \isa{P} from the conclusion to the assumptions. Then
-we apply universal introduction (\isa{allI}).
-\begin{isabelle}
-\isacommand{apply}\ (rule\ impI,\ rule\ allI)\isanewline
-\ 1.\ \isasymAnd x.\ \isasymlbrakk{\isasymforall}x.\ P\ \isasymlongrightarrow\ Q\
-x;\ P\isasymrbrakk\ \isasymLongrightarrow\ Q\ x
-\end{isabelle}
-As before, it replaces the HOL
-quantifier by a meta-level quantifier, producing a subgoal that
-binds the variable~\isa{x}. The leading bound variables
-(here \isa{x}) and the assumptions (here \isa{{\isasymforall}x.\ P\
-\isasymlongrightarrow\ Q\ x} and \isa{P}) form the \textbf{context} for the
-conclusion, here \isa{Q\ x}. Subgoals inherit the context,
-although assumptions can be added or deleted (as we saw
-earlier), while rules such as \isa{allI} add bound variables.
-
-Now, to reason from the universally quantified
-assumption, we apply the elimination rule using the \isa{drule}
-method. This rule is called \isa{spec} because it specializes a universal formula
-to a particular term.
-\begin{isabelle}
-\isacommand{apply}\ (drule\ spec)\isanewline
-\ 1.\ \isasymAnd x.\ \isasymlbrakk P;\ P\ \isasymlongrightarrow\ Q\ (?x2\
-x)\isasymrbrakk\ \isasymLongrightarrow\ Q\ x
-\end{isabelle}
-Observe how the context has changed. The quantified formula is gone,
-replaced by a new assumption derived from its body. We have
-removed the quantifier and replaced the bound variable
-by the curious term
-\isa{?x2~x}. This term is a placeholder: it may become any term that can be
-built from~\isa{x}. (Formally, \isa{?x2} is an unknown of function type, applied
-to the argument~\isa{x}.) This new assumption is an implication, so we can use
-\emph{modus ponens} on it, which concludes the proof.
-\begin{isabelle}
-\isacommand{by}\ (drule\ mp)
-\end{isabelle}
-Let us take a closer look at this last step. \emph{Modus ponens} yields
-two subgoals: one where we prove the antecedent (in this case \isa{P}) and
-one where we may assume the consequent. Both of these subgoals are proved
-by the
-\isa{assumption} method, which is implicit in the
-\isacommand{by} command. Replacing the \isacommand{by} command by
-\isa{\isacommand{apply} (drule\ mp, assumption)} would have left one last
-subgoal:
-\begin{isabelle}
-\ 1.\ \isasymAnd x.\ \isasymlbrakk P;\ Q\ (?x2\ x)\isasymrbrakk\
-\isasymLongrightarrow\ Q\ x
-\end{isabelle}
-The consequent is \isa{Q} applied to that placeholder. It may be replaced by any
-term built from~\isa{x}, and here
-it should simply be~\isa{x}. The assumption need not
-be identical to the conclusion, provided the two formulas are unifiable.%
-\index{quantifiers!universal|)}
-
-
-\subsection{The Existential Quantifier}
-
-\index{quantifiers!existential|(}%
-The concepts just presented also apply
-to the existential quantifier, whose introduction rule looks like this in
-Isabelle:
-\begin{isabelle}
-?P\ ?x\ \isasymLongrightarrow\ {\isasymexists}x.\ ?P\ x\rulenamedx{exI}
-\end{isabelle}
-If we can exhibit some $x$ such that $P(x)$ is true, then $\exists x.
-P(x)$ is also true. It is a dual of the universal elimination rule, and
-logic texts present it using the same notation for substitution.
-
-The existential
-elimination rule looks like this
-in a logic text:
-\[ \infer{Q}{\exists x.\,P & \infer*{Q}{[P]}} \]
-%
-It looks like this in Isabelle:
-\begin{isabelle}
-\isasymlbrakk{\isasymexists}x.\ ?P\ x;\ \isasymAnd x.\ ?P\ x\ \isasymLongrightarrow\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?Q\rulenamedx{exE}
-\end{isabelle}
-%
-Given an existentially quantified theorem and some
-formula $Q$ to prove, it creates a new assumption by removing the quantifier. As with
-the universal introduction rule, the textbook version imposes a proviso on the
-quantified variable, which Isabelle expresses using its meta-logic. It is
-enough to have a universal quantifier in the meta-logic; we do not need an existential
-quantifier to be built in as well.
-
-
-\begin{exercise}
-Prove the lemma
-\[ \exists x.\, P\conj Q(x)\Imp P\conj(\exists x.\, Q(x)). \]
-\emph{Hint}: the proof is similar
-to the one just above for the universal quantifier.
-\end{exercise}
-\index{quantifiers!existential|)}
-
-
-\subsection{Renaming a Bound Variable: {\tt\slshape rename_tac}}
-
-\index{assumptions!renaming|(}\index{*rename_tac (method)|(}%
-When you apply a rule such as \isa{allI}, the quantified variable
-becomes a new bound variable of the new subgoal. Isabelle tries to avoid
-changing its name, but sometimes it has to choose a new name in order to
-avoid a clash. The result may not be ideal:
-\begin{isabelle}
-\isacommand{lemma}\ "x\ <\ y\ \isasymLongrightarrow \ \isasymforall x\ y.\ P\ x\
-(f\ y)"\isanewline
-\isacommand{apply}\ (intro allI)\isanewline
-\ 1.\ \isasymAnd xa\ ya.\ x\ <\ y\ \isasymLongrightarrow \ P\ xa\ (f\ ya)
-\end{isabelle}
-%
-The names \isa{x} and \isa{y} were already in use, so the new bound variables are
-called \isa{xa} and~\isa{ya}. You can rename them by invoking \isa{rename_tac}:
-
-\begin{isabelle}
-\isacommand{apply}\ (rename_tac\ v\ w)\isanewline
-\ 1.\ \isasymAnd v\ w.\ x\ <\ y\ \isasymLongrightarrow \ P\ v\ (f\ w)
-\end{isabelle}
-Recall that \isa{rule_tac}\index{*rule_tac (method)!and renaming}
-instantiates a
-theorem with specified terms. These terms may involve the goal's bound
-variables, but beware of referring to variables
-like~\isa{xa}. A future change to your theories could change the set of names
-produced at top level, so that \isa{xa} changes to~\isa{xb} or reverts to~\isa{x}.
-It is safer to rename automatically-generated variables before mentioning them.
-
-If the subgoal has more bound variables than there are names given to
-\isa{rename_tac}, the rightmost ones are renamed.%
-\index{assumptions!renaming|)}\index{*rename_tac (method)|)}
-
-
-\subsection{Reusing an Assumption: {\tt\slshape frule}}
-\label{sec:frule}
-
-\index{assumptions!reusing|(}\index{*frule (method)|(}%
-Note that \isa{drule spec} removes the universal quantifier and --- as
-usual with elimination rules --- discards the original formula. Sometimes, a
-universal formula has to be kept so that it can be used again. Then we use a new
-method: \isa{frule}. It acts like \isa{drule} but copies rather than replaces
-the selected assumption. The \isa{f} is for \emph{forward}.
-
-In this example, going from \isa{P\ a} to \isa{P(h(h~a))}
-requires two uses of the quantified assumption, one for each~\isa{h}
-in~\isa{h(h~a)}.
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ (h\ x);
-\ P\ a\isasymrbrakk\ \isasymLongrightarrow\ P(h\ (h\ a))"
-\end{isabelle}
-%
-Examine the subgoal left by \isa{frule}:
-\begin{isabelle}
-\isacommand{apply}\ (frule\ spec)\isanewline
-\ 1.\ \isasymlbrakk{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ (h\ x);\ P\ a;\ P\ ?x\ \isasymlongrightarrow\ P\ (h\ ?x)\isasymrbrakk\ \isasymLongrightarrow\ P\ (h\ (h\ a))
-\end{isabelle}
-It is what \isa{drule} would have left except that the quantified
-assumption is still present. Next we apply \isa{mp} to the
-implication and the assumption~\isa{P\ a}:
-\begin{isabelle}
-\isacommand{apply}\ (drule\ mp,\ assumption)\isanewline
-\ 1.\ \isasymlbrakk{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ (h\ x);\ P\ a;\ P\ (h\ a)\isasymrbrakk\ \isasymLongrightarrow\ P\ (h\ (h\ a))
-\end{isabelle}
-%
-We have created the assumption \isa{P(h\ a)}, which is progress. To
-continue the proof, we apply \isa{spec} again. We shall not need it
-again, so we can use
-\isa{drule}.
-\begin{isabelle}
-\isacommand{apply}\ (drule\ spec)\isanewline
-\ 1.\ \isasymlbrakk P\ a;\ P\ (h\ a);\ P\ ?x2\
-\isasymlongrightarrow \ P\ (h\ ?x2)\isasymrbrakk \ \isasymLongrightarrow \
-P\ (h\ (h\ a))
-\end{isabelle}
-%
-The new assumption bridges the gap between \isa{P(h\ a)} and \isa{P(h(h\ a))}.
-\begin{isabelle}
-\isacommand{by}\ (drule\ mp)
-\end{isabelle}
-
-\medskip
-\emph{A final remark}. Replacing this \isacommand{by} command with
-\begin{isabelle}
-\isacommand{apply}\ (drule\ mp,\ assumption)
-\end{isabelle}
-would not work: it would add a second copy of \isa{P(h~a)} instead
-of the desired assumption, \isa{P(h(h~a))}. The \isacommand{by}
-command forces Isabelle to backtrack until it finds the correct one.
-Alternatively, we could have used the \isacommand{apply} command and bundled the
-\isa{drule mp} with \emph{two} calls of \isa{assumption}. Or, of course,
-we could have given the entire proof to \isa{auto}.%
-\index{assumptions!reusing|)}\index{*frule (method)|)}
-
-
-
-\subsection{Instantiating a Quantifier Explicitly}
-\index{quantifiers!instantiating}
-
-We can prove a theorem of the form $\exists x.\,P\, x$ by exhibiting a
-suitable term~$t$ such that $P\,t$ is true. Dually, we can use an
-assumption of the form $\forall x.\,P\, x$ to generate a new assumption $P\,t$ for
-a suitable term~$t$. In many cases,
-Isabelle makes the correct choice automatically, constructing the term by
-unification. In other cases, the required term is not obvious and we must
-specify it ourselves. Suitable methods are \isa{rule_tac}, \isa{drule_tac}
-and \isa{erule_tac}.
-
-We have seen (just above, {\S}\ref{sec:frule}) a proof of this lemma:
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk \isasymforall x.\ P\ x\
-\isasymlongrightarrow \ P\ (h\ x);\ P\ a\isasymrbrakk \
-\isasymLongrightarrow \ P(h\ (h\ a))"
-\end{isabelle}
-We had reached this subgoal:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk{\isasymforall}x.\ P\ x\ \isasymlongrightarrow\ P\ (h\
-x);\ P\ a;\ P\ (h\ a)\isasymrbrakk\ \isasymLongrightarrow\ P\ (h\ (h\ a))
-\end{isabelle}
-%
-The proof requires instantiating the quantified assumption with the
-term~\isa{h~a}.
-\begin{isabelle}
-\isacommand{apply}\ (drule_tac\ x\ =\ "h\ a"\ \isakeyword{in}\
-spec)\isanewline
-\ 1.\ \isasymlbrakk P\ a;\ P\ (h\ a);\ P\ (h\ a)\ \isasymlongrightarrow \
-P\ (h\ (h\ a))\isasymrbrakk \ \isasymLongrightarrow \ P\ (h\ (h\ a))
-\end{isabelle}
-We have forced the desired instantiation.
-
-\medskip
-Existential formulas can be instantiated too. The next example uses the
-\textbf{divides} relation\index{divides relation}
-of number theory:
-\begin{isabelle}
-?m\ dvd\ ?n\ \isasymequiv\ {\isasymexists}k.\ ?n\ =\ ?m\ *\ k
-\rulename{dvd_def}
-\end{isabelle}
-
-Let us prove that multiplication of natural numbers is monotone with
-respect to the divides relation:
-\begin{isabelle}
-\isacommand{lemma}\ mult_dvd_mono:\ "{\isasymlbrakk}i\ dvd\ m;\ j\ dvd\
-n\isasymrbrakk\ \isasymLongrightarrow\ i*j\ dvd\ (m*n\ ::\ nat)"\isanewline
-\isacommand{apply}\ (simp\ add:\ dvd_def)
-\end{isabelle}
-%
-Unfolding the definition of divides has left this subgoal:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk \isasymexists k.\ m\ =\ i\ *\ k;\ \isasymexists k.\ n\
-=\ j\ *\ k\isasymrbrakk \ \isasymLongrightarrow \ \isasymexists k.\ m\ *\
-n\ =\ i\ *\ j\ *\ k
-\end{isabelle}
-%
-Next, we eliminate the two existential quantifiers in the assumptions:
-\begin{isabelle}
-\isacommand{apply}\ (erule\ exE)\isanewline
-\ 1.\ \isasymAnd k.\ \isasymlbrakk \isasymexists k.\ n\ =\ j\ *\ k;\ m\ =\
-i\ *\ k\isasymrbrakk \ \isasymLongrightarrow \ \isasymexists k.\ m\ *\ n\
-=\ i\ *\ j\ *\ k%
-\isanewline
-\isacommand{apply}\ (erule\ exE)
-\isanewline
-\ 1.\ \isasymAnd k\ ka.\ \isasymlbrakk m\ =\ i\ *\ k;\ n\ =\ j\ *\
-ka\isasymrbrakk \ \isasymLongrightarrow \ \isasymexists k.\ m\ *\ n\ =\ i\
-*\ j\ *\ k
-\end{isabelle}
-%
-The term needed to instantiate the remaining quantifier is~\isa{k*ka}. But
-\isa{ka} is an automatically-generated name. As noted above, references to
-such variable names makes a proof less resilient to future changes. So,
-first we rename the most recent variable to~\isa{l}:
-\begin{isabelle}
-\isacommand{apply}\ (rename_tac\ l)\isanewline
-\ 1.\ \isasymAnd k\ l.\ \isasymlbrakk m\ =\ i\ *\ k;\ n\ =\ j\ *\ l\isasymrbrakk \
-\isasymLongrightarrow \ \isasymexists k.\ m\ *\ n\ =\ i\ *\ j\ *\ k%
-\end{isabelle}
-
-We instantiate the quantifier with~\isa{k*l}:
-\begin{isabelle}
-\isacommand{apply}\ (rule_tac\ x="k*l"\ \isakeyword{in}\ exI)\ \isanewline
-\ 1.\ \isasymAnd k\ ka.\ \isasymlbrakk m\ =\ i\ *\ k;\ n\ =\ j\ *\
-ka\isasymrbrakk \ \isasymLongrightarrow \ m\ *\ n\ =\ i\
-*\ j\ *\ (k\ *\ ka)
-\end{isabelle}
-%
-The rest is automatic, by arithmetic.
-\begin{isabelle}
-\isacommand{apply}\ simp\isanewline
-\isacommand{done}\isanewline
-\end{isabelle}
-
-
-
-\section{Description Operators}
-\label{sec:SOME}
-
-\index{description operators|(}%
-HOL provides two description operators.
-A \textbf{definite description} formalizes the word ``the,'' as in
-``the greatest divisior of~$n$.''
-It returns an arbitrary value unless the formula has a unique solution.
-An \textbf{indefinite description} formalizes the word ``some,'' as in
-``some member of~$S$.'' It differs from a definite description in not
-requiring the solution to be unique: it uses the axiom of choice to pick any
-solution.
-
-\begin{warn}
-Description operators can be hard to reason about. Novices
-should try to avoid them. Fortunately, descriptions are seldom required.
-\end{warn}
-
-\subsection{Definite Descriptions}
-
-\index{descriptions!definite}%
-A definite description is traditionally written $\iota x. P(x)$. It denotes
-the $x$ such that $P(x)$ is true, provided there exists a unique such~$x$;
-otherwise, it returns an arbitrary value of the expected type.
-Isabelle uses \sdx{THE} for the Greek letter~$\iota$.
-
-%(The traditional notation could be provided, but it is not legible on screen.)
-
-We reason using this rule, where \isa{a} is the unique solution:
-\begin{isabelle}
-\isasymlbrakk P\ a;\ \isasymAnd x.\ P\ x\ \isasymLongrightarrow \ x\ =\ a\isasymrbrakk \
-\isasymLongrightarrow \ (THE\ x.\ P\ x)\ =\ a%
-\rulenamedx{the_equality}
-\end{isabelle}
-For instance, we can define the
-cardinality of a finite set~$A$ to be that
-$n$ such that $A$ is in one-to-one correspondence with $\{1,\ldots,n\}$. We can then
-prove that the cardinality of the empty set is zero (since $n=0$ satisfies the
-description) and proceed to prove other facts.
-
-A more challenging example illustrates how Isabelle/HOL defines the least number
-operator, which denotes the least \isa{x} satisfying~\isa{P}:%
-\index{least number operator|see{\protect\isa{LEAST}}}
-\begin{isabelle}
-(LEAST\ x.\ P\ x)\ = (THE\ x.\ P\ x\ \isasymand \ (\isasymforall y.\
-P\ y\ \isasymlongrightarrow \ x\ \isasymle \ y))
-\end{isabelle}
-%
-Let us prove the analogue of \isa{the_equality} for \sdx{LEAST}\@.
-\begin{isabelle}
-\isacommand{theorem}\ Least_equality:\isanewline
-\ \ \ \ \ "\isasymlbrakk P\ (k::nat);\ \ \isasymforall x.\ P\ x\ \isasymlongrightarrow \ k\ \isasymle \ x\isasymrbrakk \ \isasymLongrightarrow \ (LEAST\ x.\ P\ x)\ =\ k"\isanewline
-\isacommand{apply}\ (simp\ add:\ Least_def)\isanewline
-\isanewline
-\ 1.\ \isasymlbrakk P\ k;\ \isasymforall x.\ P\ x\ \isasymlongrightarrow \ k\ \isasymle \ x\isasymrbrakk \isanewline
-\isaindent{\ 1.\ }\isasymLongrightarrow \ (THE\ x.\ P\ x\ \isasymand \ (\isasymforall y.\ P\ y\ \isasymlongrightarrow \ x\ \isasymle \ y))\ =\ k%
-\end{isabelle}
-The first step has merely unfolded the definition.
-\begin{isabelle}
-\isacommand{apply}\ (rule\ the_equality)\isanewline
-\isanewline
-\ 1.\ \isasymlbrakk P\ k;\ \isasymforall x.\ P\ x\ \isasymlongrightarrow \ k\
-\isasymle \ x\isasymrbrakk \ \isasymLongrightarrow \ P\ k\ \isasymand \
-(\isasymforall y.\ P\ y\ \isasymlongrightarrow \ k\ \isasymle \ y)\isanewline
-\ 2.\ \isasymAnd x.\ \isasymlbrakk P\ k;\ \isasymforall x.\ P\ x\ \isasymlongrightarrow \ k\ \isasymle \ x;\ P\ x\ \isasymand \ (\isasymforall y.\ P\ y\ \isasymlongrightarrow \ x\ \isasymle \ y)\isasymrbrakk \isanewline
-\ \ \ \ \ \ \ \ \isasymLongrightarrow \ x\ =\ k%
-\end{isabelle}
-As always with \isa{the_equality}, we must show existence and
-uniqueness of the claimed solution,~\isa{k}. Existence, the first
-subgoal, is trivial. Uniqueness, the second subgoal, follows by antisymmetry:
-\begin{isabelle}
-\isasymlbrakk x\ \isasymle \ y;\ y\ \isasymle \ x\isasymrbrakk \ \isasymLongrightarrow \ x\ =\ y%
-\rulename{order_antisym}
-\end{isabelle}
-The assumptions imply both \isa{k~\isasymle~x} and \isa{x~\isasymle~k}. One
-call to \isa{auto} does it all:
-\begin{isabelle}
-\isacommand{by}\ (auto\ intro:\ order_antisym)
-\end{isabelle}
-
-
-\subsection{Indefinite Descriptions}
-
-\index{Hilbert's $\varepsilon$-operator}%
-\index{descriptions!indefinite}%
-An indefinite description is traditionally written $\varepsilon x. P(x)$ and is
-known as Hilbert's $\varepsilon$-operator. It denotes
-some $x$ such that $P(x)$ is true, provided one exists.
-Isabelle uses \sdx{SOME} for the Greek letter~$\varepsilon$.
-
-Here is the definition of~\cdx{inv},\footnote{In fact, \isa{inv} is defined via a second constant \isa{inv_into}, which we ignore here.} which expresses inverses of
-functions:
-\begin{isabelle}
-inv\ f\ \isasymequiv \ \isasymlambda y.\ SOME\ x.\ f\ x\ =\ y%
-\rulename{inv_def}
-\end{isabelle}
-Using \isa{SOME} rather than \isa{THE} makes \isa{inv~f} behave well
-even if \isa{f} is not injective. As it happens, most useful theorems about
-\isa{inv} do assume the function to be injective.
-
-The inverse of \isa{f}, when applied to \isa{y}, returns some~\isa{x} such that
-\isa{f~x~=~y}. For example, we can prove \isa{inv~Suc} really is the inverse
-of the \isa{Suc} function
-\begin{isabelle}
-\isacommand{lemma}\ "inv\ Suc\ (Suc\ n)\ =\ n"\isanewline
-\isacommand{by}\ (simp\ add:\ inv_def)
-\end{isabelle}
-
-\noindent
-The proof is a one-liner: the subgoal simplifies to a degenerate application of
-\isa{SOME}, which is then erased. In detail, the left-hand side simplifies
-to \isa{SOME\ x.\ Suc\ x\ =\ Suc\ n}, then to \isa{SOME\ x.\ x\ =\ n} and
-finally to~\isa{n}.
-
-We know nothing about what
-\isa{inv~Suc} returns when applied to zero. The proof above still treats
-\isa{SOME} as a definite description, since it only reasons about
-situations in which the value is described uniquely. Indeed, \isa{SOME}
-satisfies this rule:
-\begin{isabelle}
-\isasymlbrakk P\ a;\ \isasymAnd x.\ P\ x\ \isasymLongrightarrow \ x\ =\ a\isasymrbrakk \
-\isasymLongrightarrow \ (SOME\ x.\ P\ x)\ =\ a%
-\rulenamedx{some_equality}
-\end{isabelle}
-To go further is
-tricky and requires rules such as these:
-\begin{isabelle}
-P\ x\ \isasymLongrightarrow \ P\ (SOME\ x.\ P\ x)
-\rulenamedx{someI}\isanewline
-\isasymlbrakk P\ a;\ \isasymAnd x.\ P\ x\ \isasymLongrightarrow \ Q\
-x\isasymrbrakk \ \isasymLongrightarrow \ Q\ (SOME\ x.\ P\ x)
-\rulenamedx{someI2}
-\end{isabelle}
-Rule \isa{someI} is basic: if anything satisfies \isa{P} then so does
-\hbox{\isa{SOME\ x.\ P\ x}}. The repetition of~\isa{P} in the conclusion makes it
-difficult to apply in a backward proof, so the derived rule \isa{someI2} is
-also provided.
-
-\medskip
-For example, let us prove the \rmindex{axiom of choice}:
-\begin{isabelle}
-\isacommand{theorem}\ axiom_of_choice:
-\ "(\isasymforall x.\ \isasymexists y.\ P\ x\ y)\ \isasymLongrightarrow \
-\isasymexists f.\ \isasymforall x.\ P\ x\ (f\ x)"\isanewline
-\isacommand{apply}\ (rule\ exI,\ rule\ allI)\isanewline
-
-\ 1.\ \isasymAnd x.\ \isasymforall x.\ \isasymexists y.\ P\ x\ y\
-\isasymLongrightarrow \ P\ x\ (?f\ x)
-\end{isabelle}
-%
-We have applied the introduction rules; now it is time to apply the elimination
-rules.
-
-\begin{isabelle}
-\isacommand{apply}\ (drule\ spec,\ erule\ exE)\isanewline
-
-\ 1.\ \isasymAnd x\ y.\ P\ (?x2\ x)\ y\ \isasymLongrightarrow \ P\ x\ (?f\ x)
-\end{isabelle}
-
-\noindent
-The rule \isa{someI} automatically instantiates
-\isa{f} to \hbox{\isa{\isasymlambda x.\ SOME y.\ P\ x\ y}}, which is the choice
-function. It also instantiates \isa{?x2\ x} to \isa{x}.
-\begin{isabelle}
-\isacommand{by}\ (rule\ someI)\isanewline
-\end{isabelle}
-
-\subsubsection{Historical Note}
-The original purpose of Hilbert's $\varepsilon$-operator was to express an
-existential destruction rule:
-\[ \infer{P[(\varepsilon x. P) / \, x]}{\exists x.\,P} \]
-This rule is seldom used for that purpose --- it can cause exponential
-blow-up --- but it is occasionally used as an introduction rule
-for the~$\varepsilon$-operator. Its name in HOL is \tdxbold{someI_ex}.%%
-\index{description operators|)}
-
-
-\section{Some Proofs That Fail}
-
-\index{proofs!examples of failing|(}%
-Most of the examples in this tutorial involve proving theorems. But not every
-conjecture is true, and it can be instructive to see how
-proofs fail. Here we attempt to prove a distributive law involving
-the existential quantifier and conjunction.
-\begin{isabelle}
-\isacommand{lemma}\ "({\isasymexists}x.\ P\ x)\ \isasymand\
-({\isasymexists}x.\ Q\ x)\ \isasymLongrightarrow\ {\isasymexists}x.\ P\ x\
-\isasymand\ Q\ x"
-\end{isabelle}
-The first steps are routine. We apply conjunction elimination to break
-the assumption into two existentially quantified assumptions.
-Applying existential elimination removes one of the quantifiers.
-\begin{isabelle}
-\isacommand{apply}\ (erule\ conjE)\isanewline
-\isacommand{apply}\ (erule\ exE)\isanewline
-\ 1.\ \isasymAnd x.\ \isasymlbrakk{\isasymexists}x.\ Q\ x;\ P\ x\isasymrbrakk\ \isasymLongrightarrow\ {\isasymexists}x.\ P\ x\ \isasymand\ Q\ x
-\end{isabelle}
-%
-When we remove the other quantifier, we get a different bound
-variable in the subgoal. (The name \isa{xa} is generated automatically.)
-\begin{isabelle}
-\isacommand{apply}\ (erule\ exE)\isanewline
-\ 1.\ \isasymAnd x\ xa.\ \isasymlbrakk P\ x;\ Q\ xa\isasymrbrakk\
-\isasymLongrightarrow\ {\isasymexists}x.\ P\ x\ \isasymand\ Q\ x
-\end{isabelle}
-The proviso of the existential elimination rule has forced the variables to
-differ: we can hardly expect two arbitrary values to be equal! There is
-no way to prove this subgoal. Removing the
-conclusion's existential quantifier yields two
-identical placeholders, which can become any term involving the variables \isa{x}
-and~\isa{xa}. We need one to become \isa{x}
-and the other to become~\isa{xa}, but Isabelle requires all instances of a
-placeholder to be identical.
-\begin{isabelle}
-\isacommand{apply}\ (rule\ exI)\isanewline
-\isacommand{apply}\ (rule\ conjI)\isanewline
-\ 1.\ \isasymAnd x\ xa.\ \isasymlbrakk P\ x;\ Q\ xa\isasymrbrakk\
-\isasymLongrightarrow\ P\ (?x3\ x\ xa)\isanewline
-\ 2.\ \isasymAnd x\ xa.\ \isasymlbrakk P\ x;\ Q\ xa\isasymrbrakk\ \isasymLongrightarrow\ Q\ (?x3\ x\ xa)
-\end{isabelle}
-We can prove either subgoal
-using the \isa{assumption} method. If we prove the first one, the placeholder
-changes into~\isa{x}.
-\begin{isabelle}
-\ \isacommand{apply}\ assumption\isanewline
-\ 1.\ \isasymAnd x\ xa.\ \isasymlbrakk P\ x;\ Q\ xa\isasymrbrakk\
-\isasymLongrightarrow\ Q\ x
-\end{isabelle}
-We are left with a subgoal that cannot be proved. Applying the \isa{assumption}
-method results in an error message:
-\begin{isabelle}
-*** empty result sequence -- proof command failed
-\end{isabelle}
-When interacting with Isabelle via the shell interface,
-you can abandon a proof using the \isacommand{oops} command.
-
-\medskip
-
-Here is another abortive proof, illustrating the interaction between
-bound variables and unknowns.
-If $R$ is a reflexive relation,
-is there an $x$ such that $R\,x\,y$ holds for all $y$? Let us see what happens when
-we attempt to prove it.
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymforall y.\ R\ y\ y\ \isasymLongrightarrow
-\ \isasymexists x.\ \isasymforall y.\ R\ x\ y"
-\end{isabelle}
-First, we remove the existential quantifier. The new proof state has an
-unknown, namely~\isa{?x}.
-\begin{isabelle}
-\isacommand{apply}\ (rule\ exI)\isanewline
-\ 1.\ \isasymforall y.\ R\ y\ y\ \isasymLongrightarrow \ \isasymforall y.\ R\ ?x\ y%
-\end{isabelle}
-It looks like we can just apply \isa{assumption}, but it fails. Isabelle
-refuses to substitute \isa{y}, a bound variable, for~\isa{?x}; that would be
-a bound variable capture. We can still try to finish the proof in some
-other way. We remove the universal quantifier from the conclusion, moving
-the bound variable~\isa{y} into the subgoal. But note that it is still
-bound!
-\begin{isabelle}
-\isacommand{apply}\ (rule\ allI)\isanewline
-\ 1.\ \isasymAnd y.\ \isasymforall y.\ R\ y\ y\ \isasymLongrightarrow \ R\ ?x\ y%
-\end{isabelle}
-Finally, we try to apply our reflexivity assumption. We obtain a
-new assumption whose identical placeholders may be replaced by
-any term involving~\isa{y}.
-\begin{isabelle}
-\isacommand{apply}\ (drule\ spec)\isanewline
-\ 1.\ \isasymAnd y.\ R\ (?z2\ y)\ (?z2\ y)\ \isasymLongrightarrow\ R\ ?x\ y
-\end{isabelle}
-This subgoal can only be proved by putting \isa{y} for all the placeholders,
-making the assumption and conclusion become \isa{R\ y\ y}. Isabelle can
-replace \isa{?z2~y} by \isa{y}; this involves instantiating
-\isa{?z2} to the identity function. But, just as two steps earlier,
-Isabelle refuses to substitute~\isa{y} for~\isa{?x}.
-This example is typical of how Isabelle enforces sound quantifier reasoning.
-\index{proofs!examples of failing|)}
-
-\section{Proving Theorems Using the {\tt\slshape blast} Method}
-
-\index{*blast (method)|(}%
-It is hard to prove many theorems using the methods
-described above. A proof may be hundreds of steps long. You
-may need to search among different ways of proving certain
-subgoals. Often a choice that proves one subgoal renders another
-impossible to prove. There are further complications that we have not
-discussed, concerning negation and disjunction. Isabelle's
-\textbf{classical reasoner} is a family of tools that perform such
-proofs automatically. The most important of these is the
-\isa{blast} method.
-
-In this section, we shall first see how to use the classical
-reasoner in its default mode and then how to insert additional
-rules, enabling it to work in new problem domains.
-
- We begin with examples from pure predicate logic. The following
-example is known as Andrew's challenge. Peter Andrews designed
-it to be hard to prove by automatic means.
-It is particularly hard for a resolution prover, where
-converting the nested biconditionals to
-clause form produces a combinatorial
-explosion~\cite{pelletier86}. However, the
-\isa{blast} method proves it in a fraction of a second.
-\begin{isabelle}
-\isacommand{lemma}\
-"(({\isasymexists}x.\
-{\isasymforall}y.\
-p(x){=}p(y))\
-=\
-(({\isasymexists}x.\
-q(x))=({\isasymforall}y.\
-p(y))))\
-\ \ =\ \ \ \ \isanewline
-\ \ \ \ \ \ \ \
-(({\isasymexists}x.\
-{\isasymforall}y.\
-q(x){=}q(y))\ =\ (({\isasymexists}x.\ p(x))=({\isasymforall}y.\ q(y))))"\isanewline
-\isacommand{by}\ blast
-\end{isabelle}
-The next example is a logic problem composed by Lewis Carroll.
-The \isa{blast} method finds it trivial. Moreover, it turns out
-that not all of the assumptions are necessary. We can
-experiment with variations of this formula and see which ones
-can be proved.
-\begin{isabelle}
-\isacommand{lemma}\
-"({\isasymforall}x.\
-honest(x)\ \isasymand\
-industrious(x)\ \isasymlongrightarrow\
-healthy(x))\
-\isasymand\ \ \isanewline
-\ \ \ \ \ \ \ \ \isasymnot\ ({\isasymexists}x.\
-grocer(x)\ \isasymand\
-healthy(x))\
-\isasymand\ \isanewline
-\ \ \ \ \ \ \ \ ({\isasymforall}x.\
-industrious(x)\ \isasymand\
-grocer(x)\ \isasymlongrightarrow\
-honest(x))\
-\isasymand\ \isanewline
-\ \ \ \ \ \ \ \ ({\isasymforall}x.\
-cyclist(x)\ \isasymlongrightarrow\
-industrious(x))\
-\isasymand\ \isanewline
-\ \ \ \ \ \ \ \ ({\isasymforall}x.\
-{\isasymnot}healthy(x)\ \isasymand\
-cyclist(x)\ \isasymlongrightarrow\
-{\isasymnot}honest(x))\
-\ \isanewline
-\ \ \ \ \ \ \ \ \isasymlongrightarrow\
-({\isasymforall}x.\
-grocer(x)\ \isasymlongrightarrow\
-{\isasymnot}cyclist(x))"\isanewline
-\isacommand{by}\ blast
-\end{isabelle}
-The \isa{blast} method is also effective for set theory, which is
-described in the next chapter. The formula below may look horrible, but
-the \isa{blast} method proves it in milliseconds.
-\begin{isabelle}
-\isacommand{lemma}\ "({\isasymUnion}i{\isasymin}I.\ A(i))\ \isasyminter\ ({\isasymUnion}j{\isasymin}J.\ B(j))\ =\isanewline
-\ \ \ \ \ \ \ \ ({\isasymUnion}i{\isasymin}I.\ {\isasymUnion}j{\isasymin}J.\ A(i)\ \isasyminter\ B(j))"\isanewline
-\isacommand{by}\ blast
-\end{isabelle}
-
-Few subgoals are couched purely in predicate logic and set theory.
-We can extend the scope of the classical reasoner by giving it new rules.
-Extending it effectively requires understanding the notions of
-introduction, elimination and destruction rules. Moreover, there is a
-distinction between safe and unsafe rules. A
-\textbf{safe}\indexbold{safe rules} rule is one that can be applied
-backwards without losing information; an
-\textbf{unsafe}\indexbold{unsafe rules} rule loses information, perhaps
-transforming the subgoal into one that cannot be proved. The safe/unsafe
-distinction affects the proof search: if a proof attempt fails, the
-classical reasoner backtracks to the most recent unsafe rule application
-and makes another choice.
-
-An important special case avoids all these complications. A logical
-equivalence, which in higher-order logic is an equality between
-formulas, can be given to the classical
-reasoner and simplifier by using the attribute \attrdx{iff}. You
-should do so if the right hand side of the equivalence is
-simpler than the left-hand side.
-
-For example, here is a simple fact about list concatenation.
-The result of appending two lists is empty if and only if both
-of the lists are themselves empty. Obviously, applying this equivalence
-will result in a simpler goal. When stating this lemma, we include
-the \attrdx{iff} attribute. Once we have proved the lemma, Isabelle
-will make it known to the classical reasoner (and to the simplifier).
-\begin{isabelle}
-\isacommand{lemma}\
-[iff]:\ "(xs{\isacharat}ys\ =\ [])\ =\
-(xs=[]\ \isasymand\ ys=[])"\isanewline
-\isacommand{apply}\ (induct_tac\ xs)\isanewline
-\isacommand{apply}\ (simp_all)\isanewline
-\isacommand{done}
-\end{isabelle}
-%
-This fact about multiplication is also appropriate for
-the \attrdx{iff} attribute:
-\begin{isabelle}
-(\mbox{?m}\ *\ \mbox{?n}\ =\ 0)\ =\ (\mbox{?m}\ =\ 0\ \isasymor\ \mbox{?n}\ =\ 0)
-\end{isabelle}
-A product is zero if and only if one of the factors is zero. The
-reasoning involves a disjunction. Proving new rules for
-disjunctive reasoning is hard, but translating to an actual disjunction
-works: the classical reasoner handles disjunction properly.
-
-In more detail, this is how the \attrdx{iff} attribute works. It converts
-the equivalence $P=Q$ to a pair of rules: the introduction
-rule $Q\Imp P$ and the destruction rule $P\Imp Q$. It gives both to the
-classical reasoner as safe rules, ensuring that all occurrences of $P$ in
-a subgoal are replaced by~$Q$. The simplifier performs the same
-replacement, since \isa{iff} gives $P=Q$ to the
-simplifier.
-
-Classical reasoning is different from
-simplification. Simplification is deterministic. It applies rewrite rules
-repeatedly, as long as possible, transforming a goal into another goal. Classical
-reasoning uses search and backtracking in order to prove a goal outright.%
-\index{*blast (method)|)}%
-
-
-\section{Other Classical Reasoning Methods}
-
-The \isa{blast} method is our main workhorse for proving theorems
-automatically. Other components of the classical reasoner interact
-with the simplifier. Still others perform classical reasoning
-to a limited extent, giving the user fine control over the proof.
-
-Of the latter methods, the most useful is
-\methdx{clarify}.
-It performs
-all obvious reasoning steps without splitting the goal into multiple
-parts. It does not apply unsafe rules that could render the
-goal unprovable. By performing the obvious
-steps, \isa{clarify} lays bare the difficult parts of the problem,
-where human intervention is necessary.
-
-For example, the following conjecture is false:
-\begin{isabelle}
-\isacommand{lemma}\ "({\isasymforall}x.\ P\ x)\ \isasymand\
-({\isasymexists}x.\ Q\ x)\ \isasymlongrightarrow\ ({\isasymforall}x.\ P\ x\
-\isasymand\ Q\ x)"\isanewline
-\isacommand{apply}\ clarify
-\end{isabelle}
-The \isa{blast} method would simply fail, but \isa{clarify} presents
-a subgoal that helps us see why we cannot continue the proof.
-\begin{isabelle}
-\ 1.\ \isasymAnd x\ xa.\ \isasymlbrakk{\isasymforall}x.\ P\ x;\ Q\
-xa\isasymrbrakk\ \isasymLongrightarrow\ P\ x\ \isasymand\ Q\ x
-\end{isabelle}
-The proof must fail because the assumption \isa{Q\ xa} and conclusion \isa{Q\ x}
-refer to distinct bound variables. To reach this state, \isa{clarify} applied
-the introduction rules for \isa{\isasymlongrightarrow} and \isa{\isasymforall}
-and the elimination rule for \isa{\isasymand}. It did not apply the introduction
-rule for \isa{\isasymand} because of its policy never to split goals.
-
-Also available is \methdx{clarsimp}, a method
-that interleaves \isa{clarify} and \isa{simp}. Also there is \methdx{safe},
-which like \isa{clarify} performs obvious steps but even applies those that
-split goals.
-
-The \methdx{force} method applies the classical
-reasoner and simplifier to one goal.
-Unless it can prove the goal, it fails. Contrast
-that with the \isa{auto} method, which also combines classical reasoning
-with simplification. The latter's purpose is to prove all the
-easy subgoals and parts of subgoals. Unfortunately, it can produce
-large numbers of new subgoals; also, since it proves some subgoals
-and splits others, it obscures the structure of the proof tree.
-The \isa{force} method does not have these drawbacks. Another
-difference: \isa{force} tries harder than {\isa{auto}} to prove
-its goal, so it can take much longer to terminate.
-
-Older components of the classical reasoner have largely been
-superseded by \isa{blast}, but they still have niche applications.
-Most important among these are \isa{fast} and \isa{best}. While \isa{blast}
-searches for proofs using a built-in first-order reasoner, these
-earlier methods search for proofs using standard Isabelle inference.
-That makes them slower but enables them to work in the
-presence of the more unusual features of Isabelle rules, such
-as type classes and function unknowns. For example, recall the introduction rule
-for Hilbert's $\varepsilon$-operator:
-\begin{isabelle}
-?P\ ?x\ \isasymLongrightarrow\ ?P\ (SOME\ x.\ ?P x)
-\rulename{someI}
-\end{isabelle}
-%
-The repeated occurrence of the variable \isa{?P} makes this rule tricky
-to apply. Consider this contrived example:
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk Q\ a;\ P\ a\isasymrbrakk\isanewline
-\ \ \ \ \ \ \ \ \,\isasymLongrightarrow\ P\ (SOME\ x.\ P\ x\ \isasymand\ Q\ x)\
-\isasymand\ Q\ (SOME\ x.\ P\ x\ \isasymand\ Q\ x)"\isanewline
-\isacommand{apply}\ (rule\ someI)
-\end{isabelle}
-%
-We can apply rule \isa{someI} explicitly. It yields the
-following subgoal:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk Q\ a;\ P\ a\isasymrbrakk\ \isasymLongrightarrow\ P\ ?x\
-\isasymand\ Q\ ?x%
-\end{isabelle}
-The proof from this point is trivial. Could we have
-proved the theorem with a single command? Not using \isa{blast}: it
-cannot perform the higher-order unification needed here. The
-\methdx{fast} method succeeds:
-\begin{isabelle}
-\isacommand{apply}\ (fast\ intro!:\ someI)
-\end{isabelle}
-
-The \methdx{best} method is similar to
-\isa{fast} but it uses a best-first search instead of depth-first search.
-Accordingly, it is slower but is less susceptible to divergence.
-Transitivity rules usually cause \isa{fast} to loop where \isa{best}
-can often manage.
-
-Here is a summary of the classical reasoning methods:
-\begin{itemize}
-\item \methdx{blast} works automatically and is the fastest
-
-\item \methdx{clarify} and \methdx{clarsimp} perform obvious steps without
-splitting the goal; \methdx{safe} even splits goals
-
-\item \methdx{force} uses classical reasoning and simplification to prove a goal;
- \methdx{auto} is similar but leaves what it cannot prove
-
-\item \methdx{fast} and \methdx{best} are legacy methods that work well with rules
-involving unusual features
-\end{itemize}
-A table illustrates the relationships among four of these methods.
-\begin{center}
-\begin{tabular}{r|l|l|}
- & no split & split \\ \hline
- no simp & \methdx{clarify} & \methdx{safe} \\ \hline
- simp & \methdx{clarsimp} & \methdx{auto} \\ \hline
-\end{tabular}
-\end{center}
-
-\section{Finding More Theorems}
-\label{sec:find2}
-\input{document/find2.tex}
-
-
-\section{Forward Proof: Transforming Theorems}\label{sec:forward}
-
-\index{forward proof|(}%
-Forward proof means deriving new facts from old ones. It is the
-most fundamental type of proof. Backward proof, by working from goals to
-subgoals, can help us find a difficult proof. But it is
-not always the best way of presenting the proof thus found. Forward
-proof is particularly good for reasoning from the general
-to the specific. For example, consider this distributive law for
-the greatest common divisor:
-\[ k\times\gcd(m,n) = \gcd(k\times m,k\times n)\]
-
-Putting $m=1$ we get (since $\gcd(1,n)=1$ and $k\times1=k$)
-\[ k = \gcd(k,k\times n)\]
-We have derived a new fact; if re-oriented, it might be
-useful for simplification. After re-orienting it and putting $n=1$, we
-derive another useful law:
-\[ \gcd(k,k)=k \]
-Substituting values for variables --- instantiation --- is a forward step.
-Re-orientation works by applying the symmetry of equality to
-an equation, so it too is a forward step.
-
-\subsection{Modifying a Theorem using {\tt\slshape of}, {\tt\slshape where}
- and {\tt\slshape THEN}}
-
-\label{sec:THEN}
-
-Let us reproduce our examples in Isabelle. Recall that in
-{\S}\ref{sec:fun-simplification} we declared the recursive function
-\isa{gcd}:\index{*gcd (constant)|(}
-\begin{isabelle}
-\isacommand{fun}\ gcd\ ::\ "nat\ \isasymRightarrow \ nat\ \isasymRightarrow \ nat"\ \isakeyword{where}\isanewline
-\ \ "gcd\ m\ n\ =\ (if\ n=0\ then\ m\ else\ gcd\ n\ (m\ mod\ n))"
-\end{isabelle}
-%
-From this definition, it is possible to prove the distributive law.
-That takes us to the starting point for our example.
-\begin{isabelle}
-?k\ *\ gcd\ ?m\ ?n\ =\ gcd\ (?k\ *\ ?m)\ (?k\ *\ ?n)
-\rulename{gcd_mult_distrib2}
-\end{isabelle}
-%
-The first step in our derivation is to replace \isa{?m} by~1. We instantiate the
-theorem using~\attrdx{of}, which identifies variables in order of their
-appearance from left to right. In this case, the variables are \isa{?k}, \isa{?m}
-and~\isa{?n}. So, the expression
-\hbox{\texttt{[of k 1]}} replaces \isa{?k} by~\isa{k} and \isa{?m}
-by~\isa{1}.
-\begin{isabelle}
-\isacommand{lemmas}\ gcd_mult_0\ =\ gcd_mult_distrib2\ [of\ k\ 1]
-\end{isabelle}
-%
-The keyword \commdx{lemmas} declares a new theorem, which can be derived
-from an existing one using attributes such as \isa{[of~k~1]}.
-The command
-\isa{thm gcd_mult_0}
-displays the result:
-\begin{isabelle}
-\ \ \ \ \ k\ *\ gcd\ 1\ ?n\ =\ gcd\ (k\ *\ 1)\ (k\ *\ ?n)
-\end{isabelle}
-Something is odd: \isa{k} is an ordinary variable, while \isa{?n}
-is schematic. We did not specify an instantiation
-for \isa{?n}. In its present form, the theorem does not allow
-substitution for \isa{k}. One solution is to avoid giving an instantiation for
-\isa{?k}: instead of a term we can put an underscore~(\isa{_}). For example,
-\begin{isabelle}
-\ \ \ \ \ gcd_mult_distrib2\ [of\ _\ 1]
-\end{isabelle}
-replaces \isa{?m} by~\isa{1} but leaves \isa{?k} unchanged.
-
-An equivalent solution is to use the attribute \isa{where}.
-\begin{isabelle}
-\ \ \ \ \ gcd\_mult\_distrib2\ [where\ m=1]
-\end{isabelle}
-While \isa{of} refers to
-variables by their position, \isa{where} refers to variables by name. Multiple
-instantiations are separated by~\isa{and}, as in this example:
-\begin{isabelle}
-\ \ \ \ \ gcd\_mult\_distrib2\ [where\ m=1\ and\ k=1]
-\end{isabelle}
-
-We now continue the present example with the version of \isa{gcd_mult_0}
-shown above, which has \isa{k} instead of \isa{?k}.
-Once we have replaced \isa{?m} by~1, we must next simplify
-the theorem \isa{gcd_mult_0}, performing the steps
-$\gcd(1,n)=1$ and $k\times1=k$. The \attrdx{simplified}
-attribute takes a theorem
-and returns the result of simplifying it, with respect to the default
-simplification rules:
-\begin{isabelle}
-\isacommand{lemmas}\ gcd_mult_1\ =\ gcd_mult_0\
-[simplified]%
-\end{isabelle}
-%
-Again, we display the resulting theorem:
-\begin{isabelle}
-\ \ \ \ \ k\ =\ gcd\ k\ (k\ *\ ?n)
-\end{isabelle}
-%
-To re-orient the equation requires the symmetry rule:
-\begin{isabelle}
-?s\ =\ ?t\
-\isasymLongrightarrow\ ?t\ =\
-?s%
-\rulenamedx{sym}
-\end{isabelle}
-The following declaration gives our equation to \isa{sym}:
-\begin{isabelle}
-\ \ \ \isacommand{lemmas}\ gcd_mult\ =\ gcd_mult_1\ [THEN\ sym]
-\end{isabelle}
-%
-Here is the result:
-\begin{isabelle}
-\ \ \ \ \ gcd\ k\ (k\ *\ ?n)\ =\ k%
-\end{isabelle}
-\isa{THEN~sym}\indexbold{*THEN (attribute)} gives the current theorem to the
-rule \isa{sym} and returns the resulting conclusion. The effect is to
-exchange the two operands of the equality. Typically \isa{THEN} is used
-with destruction rules. Also useful is \isa{THEN~spec}, which removes the
-quantifier from a theorem of the form $\forall x.\,P$, and \isa{THEN~mp},
-which converts the implication $P\imp Q$ into the rule
-$\vcenter{\infer{Q}{P}}$. Similar to \isa{mp} are the following two rules,
-which extract the two directions of reasoning about a boolean equivalence:
-\begin{isabelle}
-\isasymlbrakk?Q\ =\ ?P;\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P%
-\rulenamedx{iffD1}%
-\isanewline
-\isasymlbrakk?P\ =\ ?Q;\ ?Q\isasymrbrakk\ \isasymLongrightarrow\ ?P%
-\rulenamedx{iffD2}
-\end{isabelle}
-%
-Normally we would never name the intermediate theorems
-such as \isa{gcd_mult_0} and \isa{gcd_mult_1} but would combine
-the three forward steps:
-\begin{isabelle}
-\isacommand{lemmas}\ gcd_mult\ =\ gcd_mult_distrib2\ [of\ k\ 1,\ simplified,\ THEN\ sym]%
-\end{isabelle}
-The directives, or attributes, are processed from left to right. This
-declaration of \isa{gcd_mult} is equivalent to the
-previous one.
-
-Such declarations can make the proof script hard to read. Better
-is to state the new lemma explicitly and to prove it using a single
-\isa{rule} method whose operand is expressed using forward reasoning:
-\begin{isabelle}
-\isacommand{lemma}\ gcd\_mult\ [simp]:\ "gcd\ k\ (k*n)\ =\ k"\isanewline
-\isacommand{by}\ (rule\ gcd_mult_distrib2\ [of\ k\ 1,\ simplified,\ THEN\ sym])
-\end{isabelle}
-Compared with the previous proof of \isa{gcd_mult}, this
-version shows the reader what has been proved. Also, the result will be processed
-in the normal way. In particular, Isabelle generalizes over all variables: the
-resulting theorem will have {\isa{?k}} instead of {\isa{k}}.
-
-At the start of this section, we also saw a proof of $\gcd(k,k)=k$. Here
-is the Isabelle version:\index{*gcd (constant)|)}
-\begin{isabelle}
-\isacommand{lemma}\ gcd\_self\ [simp]:\ "gcd\ k\ k\ =\ k"\isanewline
-\isacommand{by}\ (rule\ gcd_mult\ [of\ k\ 1,\ simplified])
-\end{isabelle}
-
-\begin{warn}
-To give~\isa{of} a nonatomic term, enclose it in quotation marks, as in
-\isa{[of "k*m"]}. The term must not contain unknowns: an
-attribute such as
-\isa{[of "?k*m"]} will be rejected.
-\end{warn}
-
-%Answer is now included in that section! Is a modified version of this
-% exercise worth including? E.g. find a difference between the two ways
-% of substituting.
-%\begin{exercise}
-%In {\S}\ref{sec:subst} the method \isa{subst\ mult_commute} was applied. How
-%can we achieve the same effect using \isa{THEN} with the rule \isa{ssubst}?
-%% answer rule (mult_commute [THEN ssubst])
-%\end{exercise}
-
-\subsection{Modifying a Theorem using {\tt\slshape OF}}
-
-\index{*OF (attribute)|(}%
-Recall that \isa{of} generates an instance of a
-rule by specifying values for its variables. Analogous is \isa{OF}, which
-generates an instance of a rule by specifying facts for its premises.
-
-We again need the divides relation\index{divides relation} of number theory, which
-as we recall is defined by
-\begin{isabelle}
-?m\ dvd\ ?n\ \isasymequiv\ {\isasymexists}k.\ ?n\ =\ ?m\ *\ k
-\rulename{dvd_def}
-\end{isabelle}
-%
-Suppose, for example, that we have proved the following rule.
-It states that if $k$ and $n$ are relatively prime
-and if $k$ divides $m\times n$ then $k$ divides $m$.
-\begin{isabelle}
-\isasymlbrakk gcd ?k ?n {=} 1;\ ?k\ dvd\ ?m * ?n\isasymrbrakk\
-\isasymLongrightarrow\ ?k\ dvd\ ?m
-\rulename{relprime_dvd_mult}
-\end{isabelle}
-We can use \isa{OF} to create an instance of this rule.
-First, we
-prove an instance of its first premise:
-\begin{isabelle}
-\isacommand{lemma}\ relprime\_20\_81:\ "gcd\ 20\ 81\ =\ 1"\isanewline
-\isacommand{by}\ (simp\ add:\ gcd.simps)
-\end{isabelle}
-We have evaluated an application of the \isa{gcd} function by
-simplification. Expression evaluation involving recursive functions is not
-guaranteed to terminate, and it can be slow; Isabelle
-performs arithmetic by rewriting symbolic bit strings. Here,
-however, the simplification takes less than one second. We can
-give this new lemma to \isa{OF}. The expression
-\begin{isabelle}
-\ \ \ \ \ relprime_dvd_mult [OF relprime_20_81]
-\end{isabelle}
-yields the theorem
-\begin{isabelle}
-\ \ \ \ \ 20\ dvd\ (?m\ *\ 81)\ \isasymLongrightarrow\ 20\ dvd\ ?m%
-\end{isabelle}
-%
-\isa{OF} takes any number of operands. Consider
-the following facts about the divides relation:
-\begin{isabelle}
-\isasymlbrakk?k\ dvd\ ?m;\
-?k\ dvd\ ?n\isasymrbrakk\
-\isasymLongrightarrow\ ?k\ dvd\
-?m\ +\ ?n
-\rulename{dvd_add}\isanewline
-?m\ dvd\ ?m%
-\rulename{dvd_refl}
-\end{isabelle}
-Let us supply \isa{dvd_refl} for each of the premises of \isa{dvd_add}:
-\begin{isabelle}
-\ \ \ \ \ dvd_add [OF dvd_refl dvd_refl]
-\end{isabelle}
-Here is the theorem that we have expressed:
-\begin{isabelle}
-\ \ \ \ \ ?k\ dvd\ (?k\ +\ ?k)
-\end{isabelle}
-As with \isa{of}, we can use the \isa{_} symbol to leave some positions
-unspecified:
-\begin{isabelle}
-\ \ \ \ \ dvd_add [OF _ dvd_refl]
-\end{isabelle}
-The result is
-\begin{isabelle}
-\ \ \ \ \ ?k\ dvd\ ?m\ \isasymLongrightarrow\ ?k\ dvd\ ?m\ +\ ?k
-\end{isabelle}
-
-You may have noticed that \isa{THEN} and \isa{OF} are based on
-the same idea, namely to combine two rules. They differ in the
-order of the combination and thus in their effect. We use \isa{THEN}
-typically with a destruction rule to extract a subformula of the current
-theorem. We use \isa{OF} with a list of facts to generate an instance of
-the current theorem.%
-\index{*OF (attribute)|)}
-
-Here is a summary of some primitives for forward reasoning:
-\begin{itemize}
-\item \attrdx{of} instantiates the variables of a rule to a list of terms
-\item \attrdx{OF} applies a rule to a list of theorems
-\item \attrdx{THEN} gives a theorem to a named rule and returns the
-conclusion
-%\item \attrdx{rule_format} puts a theorem into standard form
-% by removing \isa{\isasymlongrightarrow} and~\isa{\isasymforall}
-\item \attrdx{simplified} applies the simplifier to a theorem
-\item \isacommand{lemmas} assigns a name to the theorem produced by the
-attributes above
-\end{itemize}
-
-
-\section{Forward Reasoning in a Backward Proof}
-
-We have seen that the forward proof directives work well within a backward
-proof. There are many ways to achieve a forward style using our existing
-proof methods. We shall also meet some new methods that perform forward
-reasoning.
-
-The methods \isa{drule}, \isa{frule}, \isa{drule_tac}, etc.,
-reason forward from a subgoal. We have seen them already, using rules such as
-\isa{mp} and
-\isa{spec} to operate on formulae. They can also operate on terms, using rules
-such as these:
-\begin{isabelle}
-x\ =\ y\ \isasymLongrightarrow \ f\ x\ =\ f\ y%
-\rulenamedx{arg_cong}\isanewline
-i\ \isasymle \ j\ \isasymLongrightarrow \ i\ *\ k\ \isasymle \ j\ *\ k%
-\rulename{mult_le_mono1}
-\end{isabelle}
-
-For example, let us prove a fact about divisibility in the natural numbers:
-\begin{isabelle}
-\isacommand{lemma}\ "2\ \isasymle \ u\ \isasymLongrightarrow \ u*m\ \isasymnoteq
-\ Suc(u*n)"\isanewline
-\isacommand{apply}\ (intro\ notI)\isanewline
-\ 1.\ \isasymlbrakk 2\ \isasymle \ u;\ u\ *\ m\ =\ Suc\ (u\ *\ n)\isasymrbrakk \ \isasymLongrightarrow \ False%
-\end{isabelle}
-%
-The key step is to apply the function \ldots\isa{mod\ u} to both sides of the
-equation
-\isa{u*m\ =\ Suc(u*n)}:\index{*drule_tac (method)}
-\begin{isabelle}
-\isacommand{apply}\ (drule_tac\ f="\isasymlambda x.\ x\ mod\ u"\ \isakeyword{in}\
-arg_cong)\isanewline
-\ 1.\ \isasymlbrakk 2\ \isasymle \ u;\ u\ *\ m\ mod\ u\ =\ Suc\ (u\ *\ n)\ mod\
-u\isasymrbrakk \ \isasymLongrightarrow \ False
-\end{isabelle}
-%
-Simplification reduces the left side to 0 and the right side to~1, yielding the
-required contradiction.
-\begin{isabelle}
-\isacommand{apply}\ (simp\ add:\ mod_Suc)\isanewline
-\isacommand{done}
-\end{isabelle}
-
-Our proof has used a fact about remainder:
-\begin{isabelle}
-Suc\ m\ mod\ n\ =\isanewline
-(if\ Suc\ (m\ mod\ n)\ =\ n\ then\ 0\ else\ Suc\ (m\ mod\ n))
-\rulename{mod_Suc}
-\end{isabelle}
-
-\subsection{The Method {\tt\slshape insert}}
-
-\index{*insert (method)|(}%
-The \isa{insert} method
-inserts a given theorem as a new assumption of all subgoals. This
-already is a forward step; moreover, we may (as always when using a
-theorem) apply
-\isa{of}, \isa{THEN} and other directives. The new assumption can then
-be used to help prove the subgoals.
-
-For example, consider this theorem about the divides relation. The first
-proof step inserts the distributive law for
-\isa{gcd}. We specify its variables as shown.
-\begin{isabelle}
-\isacommand{lemma}\ relprime\_dvd\_mult:\ \isanewline
-\ \ \ \ \ \ "\isasymlbrakk \ gcd\ k\ n\ =\ 1;\ k\ dvd\ m*n\ \isasymrbrakk \ \isasymLongrightarrow \ k\ dvd\ m"\isanewline
-\isacommand{apply}\ (insert\ gcd_mult_distrib2\ [of\ m\ k\ n])
-\end{isabelle}
-In the resulting subgoal, note how the equation has been
-inserted:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk gcd\ k\ n\ =\ 1;\ k\ dvd\ m\ *\ n;\ m\ *\ gcd\ k\ n\ =\ gcd\ (m\ *\ k)\ (m\ *\ n)\isasymrbrakk \isanewline
-\isaindent{\ 1.\ }\isasymLongrightarrow \ k\ dvd\ m%
-\end{isabelle}
-The next proof step utilizes the assumption \isa{gcd\ k\ n\ =\ 1}
-(note that \isa{Suc\ 0} is another expression for 1):
-\begin{isabelle}
-\isacommand{apply}(simp)\isanewline
-\ 1.\ \isasymlbrakk gcd\ k\ n\ =\ Suc\ 0;\ k\ dvd\ m\ *\ n;\ m\ =\ gcd\ (m\ *\ k)\ (m\ *\ n)\isasymrbrakk \isanewline
-\isaindent{\ 1.\ }\isasymLongrightarrow \ k\ dvd\ m%
-\end{isabelle}
-Simplification has yielded an equation for~\isa{m}. The rest of the proof
-is omitted.
-
-\medskip
-Here is another demonstration of \isa{insert}. Division and
-remainder obey a well-known law:
-\begin{isabelle}
-(?m\ div\ ?n)\ *\ ?n\ +\ ?m\ mod\ ?n\ =\ ?m
-\rulename{mod_div_equality}
-\end{isabelle}
-
-We refer to this law explicitly in the following proof:
-\begin{isabelle}
-\isacommand{lemma}\ div_mult_self_is_m:\ \isanewline
-\ \ \ \ \ \ "0{\isacharless}n\ \isasymLongrightarrow\ (m*n)\ div\ n\ =\
-(m::nat)"\isanewline
-\isacommand{apply}\ (insert\ mod_div_equality\ [of\ "m*n"\ n])\isanewline
-\isacommand{apply}\ (simp)\isanewline
-\isacommand{done}
-\end{isabelle}
-The first step inserts the law, specifying \isa{m*n} and
-\isa{n} for its variables. Notice that non-trivial expressions must be
-enclosed in quotation marks. Here is the resulting
-subgoal, with its new assumption:
-\begin{isabelle}
-%0\ \isacharless\ n\ \isasymLongrightarrow\ (m\
-%*\ n)\ div\ n\ =\ m\isanewline
-\ 1.\ \isasymlbrakk0\ \isacharless\
-n;\ \ (m\ *\ n)\ div\ n\ *\ n\ +\ (m\ *\ n)\ mod\ n\
-=\ m\ *\ n\isasymrbrakk\isanewline
-\ \ \ \ \isasymLongrightarrow\ (m\ *\ n)\ div\ n\
-=\ m
-\end{isabelle}
-Simplification reduces \isa{(m\ *\ n)\ mod\ n} to zero.
-Then it cancels the factor~\isa{n} on both
-sides of the equation \isa{(m\ *\ n)\ div\ n\ *\ n\ =\ m\ *\ n}, proving the
-theorem.
-
-\begin{warn}
-Any unknowns in the theorem given to \methdx{insert} will be universally
-quantified in the new assumption.
-\end{warn}%
-\index{*insert (method)|)}
-
-\subsection{The Method {\tt\slshape subgoal_tac}}
-
-\index{*subgoal_tac (method)}%
-A related method is \isa{subgoal_tac}, but instead
-of inserting a theorem as an assumption, it inserts an arbitrary formula.
-This formula must be proved later as a separate subgoal. The
-idea is to claim that the formula holds on the basis of the current
-assumptions, to use this claim to complete the proof, and finally
-to justify the claim. It gives the proof
-some structure. If you find yourself generating a complex assumption by a
-long series of forward steps, consider using \isa{subgoal_tac} instead: you can
-state the formula you are aiming for, and perhaps prove it automatically.
-
-Look at the following example.
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk(z::int)\ <\ 37;\ 66\ <\ 2*z;\ z*z\
-\isasymnoteq\ 1225;\ Q(34);\ Q(36)\isasymrbrakk\isanewline
-\ \ \ \ \ \ \ \ \,\isasymLongrightarrow\ Q(z)"\isanewline
-\isacommand{apply}\ (subgoal_tac\ "z\ =\ 34\ \isasymor\ z\ =\
-36")\isanewline
-\isacommand{apply}\ blast\isanewline
-\isacommand{apply}\ (subgoal_tac\ "z\ \isasymnoteq\ 35")\isanewline
-\isacommand{apply}\ arith\isanewline
-\isacommand{apply}\ force\isanewline
-\isacommand{done}
-\end{isabelle}
-The first assumption tells us
-that \isa{z} is no greater than~36. The second tells us that \isa{z}
-is at least~34. The third assumption tells us that \isa{z} cannot be 35, since
-$35\times35=1225$. So \isa{z} is either 34 or~36, and since \isa{Q} holds for
-both of those values, we have the conclusion.
-
-The Isabelle proof closely follows this reasoning. The first
-step is to claim that \isa{z} is either 34 or 36. The resulting proof
-state gives us two subgoals:
-\begin{isabelle}
-%\isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\ 1225;\
-%Q\ 34;\ Q\ 36\isasymrbrakk\ \isasymLongrightarrow\ Q\ z\isanewline
-\ 1.\ \isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\ 1225;\ Q\ 34;\ Q\ 36;\isanewline
-\ \ \ \ \ z\ =\ 34\ \isasymor\ z\ =\ 36\isasymrbrakk\isanewline
-\ \ \ \ \isasymLongrightarrow\ Q\ z\isanewline
-\ 2.\ \isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\ 1225;\ Q\ 34;\ Q\ 36\isasymrbrakk\isanewline
-\ \ \ \ \isasymLongrightarrow\ z\ =\ 34\ \isasymor\ z\ =\ 36
-\end{isabelle}
-The first subgoal is trivial (\isa{blast}), but for the second Isabelle needs help to eliminate
-the case
-\isa{z}=35. The second invocation of {\isa{subgoal_tac}} leaves two
-subgoals:
-\begin{isabelle}
-\ 1.\ \isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\
-1225;\ Q\ 34;\ Q\ 36;\isanewline
-\ \ \ \ \ z\ \isasymnoteq\ 35\isasymrbrakk\isanewline
-\ \ \ \ \isasymLongrightarrow\ z\ =\ 34\ \isasymor\ z\ =\ 36\isanewline
-\ 2.\ \isasymlbrakk z\ <\ 37;\ 66\ <\ 2\ *\ z;\ z\ *\ z\ \isasymnoteq\ 1225;\ Q\ 34;\ Q\ 36\isasymrbrakk\isanewline
-\ \ \ \ \isasymLongrightarrow\ z\ \isasymnoteq\ 35
-\end{isabelle}
-
-Assuming that \isa{z} is not 35, the first subgoal follows by linear arithmetic
-(\isa{arith}). For the second subgoal we apply the method \isa{force},
-which proceeds by assuming that \isa{z}=35 and arriving at a contradiction.
-
-
-\medskip
-Summary of these methods:
-\begin{itemize}
-\item \methdx{insert} adds a theorem as a new assumption
-\item \methdx{subgoal_tac} adds a formula as a new assumption and leaves the
-subgoal of proving that formula
-\end{itemize}
-\index{forward proof|)}
-
-
-\section{Managing Large Proofs}
-
-Naturally you should try to divide proofs into manageable parts. Look for lemmas
-that can be proved separately. Sometimes you will observe that they are
-instances of much simpler facts. On other occasions, no lemmas suggest themselves
-and you are forced to cope with a long proof involving many subgoals.
-
-\subsection{Tacticals, or Control Structures}
-
-\index{tacticals|(}%
-If the proof is long, perhaps it at least has some regularity. Then you can
-express it more concisely using \textbf{tacticals}, which provide control
-structures. Here is a proof (it would be a one-liner using
-\isa{blast}, but forget that) that contains a series of repeated
-commands:
-%
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk P\isasymlongrightarrow Q;\
-Q\isasymlongrightarrow R;\ R\isasymlongrightarrow S;\ P\isasymrbrakk \
-\isasymLongrightarrow \ S"\isanewline
-\isacommand{apply}\ (drule\ mp,\ assumption)\isanewline
-\isacommand{apply}\ (drule\ mp,\ assumption)\isanewline
-\isacommand{apply}\ (drule\ mp,\ assumption)\isanewline
-\isacommand{apply}\ (assumption)\isanewline
-\isacommand{done}
-\end{isabelle}
-%
-Each of the three identical commands finds an implication and proves its
-antecedent by assumption. The first one finds \isa{P\isasymlongrightarrow Q} and
-\isa{P}, concluding~\isa{Q}; the second one concludes~\isa{R} and the third one
-concludes~\isa{S}. The final step matches the assumption \isa{S} with the goal to
-be proved.
-
-Suffixing a method with a plus sign~(\isa+)\index{*"+ (tactical)}
-expresses one or more repetitions:
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk P\isasymlongrightarrow Q;\ Q\isasymlongrightarrow R;\ R\isasymlongrightarrow S;\ P\isasymrbrakk \ \isasymLongrightarrow \ S"\isanewline
-\isacommand{by}\ (drule\ mp,\ assumption)+
-\end{isabelle}
-%
-Using \isacommand{by} takes care of the final use of \isa{assumption}. The new
-proof is more concise. It is also more general: the repetitive method works
-for a chain of implications having any length, not just three.
-
-Choice is another control structure. Separating two methods by a vertical
-% we must use ?? rather than "| as the sorting item because somehow the presence
-% of | (even quoted) stops hyperref from putting |hyperpage at the end of the index
-% entry.
-bar~(\isa|)\index{??@\texttt{"|} (tactical)} gives the effect of applying the
-first method, and if that fails, trying the second. It can be combined with
-repetition, when the choice must be made over and over again. Here is a chain of
-implications in which most of the antecedents are proved by assumption, but one is
-proved by arithmetic:
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk Q\isasymlongrightarrow R;\ P\isasymlongrightarrow Q;\ x<5\isasymlongrightarrow P;\
-Suc\ x\ <\ 5\isasymrbrakk \ \isasymLongrightarrow \ R"\ \isanewline
-\isacommand{by}\ (drule\ mp,\ (assumption|arith))+
-\end{isabelle}
-The \isa{arith}
-method can prove $x<5$ from $x+1<5$, but it cannot duplicate the effect of
-\isa{assumption}. Therefore, we combine these methods using the choice
-operator.
-
-A postfixed question mark~(\isa?)\index{*"? (tactical)} expresses zero or one
-repetitions of a method. It can also be viewed as the choice between executing a
-method and doing nothing. It is useless at top level but can be valuable
-within other control structures; for example,
-\isa{($m$+)?} performs zero or more repetitions of method~$m$.%
-\index{tacticals|)}
-
-
-\subsection{Subgoal Numbering}
-
-Another problem in large proofs is contending with huge
-subgoals or many subgoals. Induction can produce a proof state that looks
-like this:
-\begin{isabelle}
-\ 1.\ bigsubgoal1\isanewline
-\ 2.\ bigsubgoal2\isanewline
-\ 3.\ bigsubgoal3\isanewline
-\ 4.\ bigsubgoal4\isanewline
-\ 5.\ bigsubgoal5\isanewline
-\ 6.\ bigsubgoal6
-\end{isabelle}
-If each \isa{bigsubgoal} is 15 lines or so, the proof state will be too big to
-scroll through. By default, Isabelle displays at most 10 subgoals. The
-\commdx{pr} command lets you change this limit:
-\begin{isabelle}
-\isacommand{pr}\ 2\isanewline
-\ 1.\ bigsubgoal1\isanewline
-\ 2.\ bigsubgoal2\isanewline
-A total of 6 subgoals...
-\end{isabelle}
-
-\medskip
-All methods apply to the first subgoal.
-Sometimes, not only in a large proof, you may want to focus on some other
-subgoal. Then you should try the commands \isacommand{defer} or \isacommand{prefer}.
-
-In the following example, the first subgoal looks hard, while the others
-look as if \isa{blast} alone could prove them:
-\begin{isabelle}
-\ 1.\ hard\isanewline
-\ 2.\ \isasymnot \ \isasymnot \ P\ \isasymLongrightarrow \ P\isanewline
-\ 3.\ Q\ \isasymLongrightarrow \ Q%
-\end{isabelle}
-%
-The \commdx{defer} command moves the first subgoal into the last position.
-\begin{isabelle}
-\isacommand{defer}\ 1\isanewline
-\ 1.\ \isasymnot \ \isasymnot \ P\ \isasymLongrightarrow \ P\isanewline
-\ 2.\ Q\ \isasymLongrightarrow \ Q\isanewline
-\ 3.\ hard%
-\end{isabelle}
-%
-Now we apply \isa{blast} repeatedly to the easy subgoals:
-\begin{isabelle}
-\isacommand{apply}\ blast+\isanewline
-\ 1.\ hard%
-\end{isabelle}
-Using \isacommand{defer}, we have cleared away the trivial parts of the proof so
-that we can devote attention to the difficult part.
-
-\medskip
-The \commdx{prefer} command moves the specified subgoal into the
-first position. For example, if you suspect that one of your subgoals is
-invalid (not a theorem), then you should investigate that subgoal first. If it
-cannot be proved, then there is no point in proving the other subgoals.
-\begin{isabelle}
-\ 1.\ ok1\isanewline
-\ 2.\ ok2\isanewline
-\ 3.\ doubtful%
-\end{isabelle}
-%
-We decide to work on the third subgoal.
-\begin{isabelle}
-\isacommand{prefer}\ 3\isanewline
-\ 1.\ doubtful\isanewline
-\ 2.\ ok1\isanewline
-\ 3.\ ok2
-\end{isabelle}
-If we manage to prove \isa{doubtful}, then we can work on the other
-subgoals, confident that we are not wasting our time. Finally we revise the
-proof script to remove the \isacommand{prefer} command, since we needed it only to
-focus our exploration. The previous example is different: its use of
-\isacommand{defer} stops trivial subgoals from cluttering the rest of the
-proof. Even there, we should consider proving \isa{hard} as a preliminary
-lemma. Always seek ways to streamline your proofs.
-
-
-\medskip
-Summary:
-\begin{itemize}
-\item the control structures \isa+, \isa? and \isa| help express complicated proofs
-\item the \isacommand{pr} command can limit the number of subgoals to display
-\item the \isacommand{defer} and \isacommand{prefer} commands move a
-subgoal to the last or first position
-\end{itemize}
-
-\begin{exercise}
-Explain the use of \isa? and \isa+ in this proof.
-\begin{isabelle}
-\isacommand{lemma}\ "\isasymlbrakk P\isasymand Q\isasymlongrightarrow R;\ P\isasymlongrightarrow Q;\ P\isasymrbrakk \ \isasymLongrightarrow \ R"\isanewline
-\isacommand{by}\ (drule\ mp,\ (intro conjI)?,\ assumption+)+
-\end{isabelle}
-\end{exercise}
-
-
-
-\section{Proving the Correctness of Euclid's Algorithm}
-\label{sec:proving-euclid}
-
-\index{Euclid's algorithm|(}\index{*gcd (constant)|(}\index{divides relation|(}%
-A brief development will demonstrate the techniques of this chapter,
-including \isa{blast} applied with additional rules. We shall also see
-\isa{case_tac} used to perform a Boolean case split.
-
-Let us prove that \isa{gcd} computes the greatest common
-divisor of its two arguments.
-%
-We use induction: \isa{gcd.induct} is the
-induction rule returned by \isa{fun}. We simplify using
-rules proved in {\S}\ref{sec:fun-simplification}, since rewriting by the
-definition of \isa{gcd} can loop.
-\begin{isabelle}
-\isacommand{lemma}\ gcd\_dvd\_both:\ "(gcd\ m\ n\ dvd\ m)\ \isasymand \ (gcd\ m\ n\ dvd\ n)"
-\end{isabelle}
-The induction formula must be a conjunction. In the
-inductive step, each conjunct establishes the other.
-\begin{isabelle}
-\ 1.\ \isasymAnd m\ n.\ (n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ (}gcd\ n\ (m\ mod\ n)\ dvd\ n\ \isasymand \isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ (}gcd\ n\ (m\ mod\ n)\ dvd\ m\ mod\ n)\ \isasymLongrightarrow \isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ }gcd\ m\ n\ dvd\ m\ \isasymand \ gcd\ m\ n\ dvd\ n%
-\end{isabelle}
-
-The conditional induction hypothesis suggests doing a case
-analysis on \isa{n=0}. We apply \methdx{case_tac} with type
-\isa{bool} --- and not with a datatype, as we have done until now. Since
-\isa{nat} is a datatype, we could have written
-\isa{case_tac~n} instead of \isa{case_tac~"n=0"}. However, the definition
-of \isa{gcd} makes a Boolean decision:
-\begin{isabelle}
-\ \ \ \ "gcd\ m\ n\ =\ (if\ n=0\ then\ m\ else\ gcd\ n\ (m\ mod\ n))"
-\end{isabelle}
-Proofs about a function frequently follow the function's definition, so we perform
-case analysis over the same formula.
-\begin{isabelle}
-\isacommand{apply}\ (case_tac\ "n=0")\isanewline
-\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk }gcd\ n\ (m\ mod\ n)\ dvd\ n\ \isasymand \ gcd\ n\ (m\ mod\ n)\ dvd\ m\ mod\ n;\isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ \ }n\ =\ 0\isasymrbrakk \isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ gcd\ m\ n\ dvd\ m\ \isasymand \ gcd\ m\ n\ dvd\ n\isanewline
-\ 2.\ \isasymAnd m\ n.\ \isasymlbrakk n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline
-\isaindent{\ 2.\ \isasymAnd m\ n.\ \isasymlbrakk }gcd\ n\ (m\ mod\ n)\ dvd\ n\ \isasymand \ gcd\ n\ (m\ mod\ n)\ dvd\ m\ mod\ n;\isanewline
-\isaindent{\ 2.\ \isasymAnd m\ n.\ \ }n\ \isasymnoteq \ 0\isasymrbrakk \isanewline
-\isaindent{\ 2.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ gcd\ m\ n\ dvd\ m\ \isasymand \ gcd\ m\ n\ dvd\ n%
-\end{isabelle}
-%
-Simplification leaves one subgoal:
-\begin{isabelle}
-\isacommand{apply}\ (simp_all)\isanewline
-\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk gcd\ n\ (m\ mod\ n)\ dvd\ n\ \isasymand \ gcd\ n\ (m\ mod\ n)\ dvd\ m\ mod\ n;\isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ \ }0\ <\ n\isasymrbrakk \isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ gcd\ n\ (m\ mod\ n)\ dvd\ m%
-\end{isabelle}
-%
-Here, we can use \isa{blast}.
-One of the assumptions, the induction hypothesis, is a conjunction.
-The two divides relationships it asserts are enough to prove
-the conclusion, for we have the following theorem at our disposal:
-\begin{isabelle}
-\isasymlbrakk?k\ dvd\ (?m\ mod\ ?n){;}\ ?k\ dvd\ ?n\isasymrbrakk\ \isasymLongrightarrow\ ?k\ dvd\ ?m%
-\rulename{dvd_mod_imp_dvd}
-\end{isabelle}
-%
-This theorem can be applied in various ways. As an introduction rule, it
-would cause backward chaining from the conclusion (namely
-\isa{?k~dvd~?m}) to the two premises, which
-also involve the divides relation. This process does not look promising
-and could easily loop. More sensible is to apply the rule in the forward
-direction; each step would eliminate an occurrence of the \isa{mod} symbol, so the
-process must terminate.
-\begin{isabelle}
-\isacommand{apply}\ (blast\ dest:\ dvd_mod_imp_dvd)\isanewline
-\isacommand{done}
-\end{isabelle}
-Attaching the \attrdx{dest} attribute to \isa{dvd_mod_imp_dvd} tells
-\isa{blast} to use it as destruction rule; that is, in the forward direction.
-
-\medskip
-We have proved a conjunction. Now, let us give names to each of the
-two halves:
-\begin{isabelle}
-\isacommand{lemmas}\ gcd_dvd1\ [iff]\ =\ gcd_dvd_both\ [THEN\ conjunct1]\isanewline
-\isacommand{lemmas}\ gcd_dvd2\ [iff]\ =\ gcd_dvd_both\ [THEN\ conjunct2]%
-\end{isabelle}
-Here we see \commdx{lemmas}
-used with the \attrdx{iff} attribute, which supplies the new theorems to the
-classical reasoner and the simplifier. Recall that \attrdx{THEN} is
-frequently used with destruction rules; \isa{THEN conjunct1} extracts the
-first half of a conjunctive theorem. Given \isa{gcd_dvd_both} it yields
-\begin{isabelle}
-\ \ \ \ \ gcd\ ?m1\ ?n1\ dvd\ ?m1
-\end{isabelle}
-The variable names \isa{?m1} and \isa{?n1} arise because
-Isabelle renames schematic variables to prevent
-clashes. The second \isacommand{lemmas} declaration yields
-\begin{isabelle}
-\ \ \ \ \ gcd\ ?m1\ ?n1\ dvd\ ?n1
-\end{isabelle}
-
-\medskip
-To complete the verification of the \isa{gcd} function, we must
-prove that it returns the greatest of all the common divisors
-of its arguments. The proof is by induction, case analysis and simplification.
-\begin{isabelle}
-\isacommand{lemma}\ gcd\_greatest\ [rule\_format]:\isanewline
-\ \ \ \ \ \ "k\ dvd\ m\ \isasymlongrightarrow \ k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ m\ n"
-\end{isabelle}
-%
-The goal is expressed using HOL implication,
-\isa{\isasymlongrightarrow}, because the induction affects the two
-preconditions. The directive \isa{rule_format} tells Isabelle to replace
-each \isa{\isasymlongrightarrow} by \isa{\isasymLongrightarrow} before
-storing the eventual theorem. This directive can also remove outer
-universal quantifiers, converting the theorem into the usual format for
-inference rules. It can replace any series of applications of
-\isa{THEN} to the rules \isa{mp} and \isa{spec}. We did not have to
-write this:
-\begin{isabelle}
-\isacommand{lemma}\ gcd_greatest\
-[THEN mp, THEN mp]:\isanewline
-\ \ \ \ \ \ "k\ dvd\ m\ \isasymlongrightarrow \ k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ m\ n"
-\end{isabelle}
-
-Because we are again reasoning about \isa{gcd}, we perform the same
-induction and case analysis as in the previous proof:
-\begingroup\samepage
-\begin{isabelle}
-\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ \isasymlbrakk }k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ m\ mod\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ n\ (m\ mod\ n);\isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ \ }n\ =\ 0\isasymrbrakk \isanewline
-\isaindent{\ 1.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ k\ dvd\ m\ \isasymlongrightarrow \ k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ m\ n\isanewline
-\ 2.\ \isasymAnd m\ n.\ \isasymlbrakk n\ \isasymnoteq \ 0\ \isasymLongrightarrow \isanewline
-\isaindent{\ 2.\ \isasymAnd m\ n.\ \isasymlbrakk }k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ m\ mod\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ n\ (m\ mod\ n);\isanewline
-\isaindent{\ 2.\ \isasymAnd m\ n.\ \ }n\ \isasymnoteq \ 0\isasymrbrakk \isanewline
-\isaindent{\ 2.\ \isasymAnd m\ n.\ }\isasymLongrightarrow \ k\ dvd\ m\ \isasymlongrightarrow \ k\ dvd\ n\ \isasymlongrightarrow \ k\ dvd\ gcd\ m\ n%
-\end{isabelle}
-\endgroup
-
-\noindent Simplification proves both subgoals.
-\begin{isabelle}
-\isacommand{apply}\ (simp_all\ add:\ dvd_mod)\isanewline
-\isacommand{done}
-\end{isabelle}
-In the first, where \isa{n=0}, the implication becomes trivial: \isa{k\ dvd\
-gcd\ m\ n} goes to~\isa{k\ dvd\ m}. The second subgoal is proved by
-an unfolding of \isa{gcd}, using this rule about divides:
-\begin{isabelle}
-\isasymlbrakk ?f\ dvd\ ?m;\ ?f\ dvd\ ?n\isasymrbrakk \
-\isasymLongrightarrow \ ?f\ dvd\ ?m\ mod\ ?n%
-\rulename{dvd_mod}
-\end{isabelle}
-
-
-\medskip
-The facts proved above can be summarized as a single logical
-equivalence. This step gives us a chance to see another application
-of \isa{blast}.
-\begin{isabelle}
-\isacommand{theorem}\ gcd\_greatest\_iff\ [iff]:\ \isanewline
-\ \ \ \ \ \ \ \ "(k\ dvd\ gcd\ m\ n)\ =\ (k\ dvd\ m\ \isasymand \ k\ dvd\ n)"\isanewline
-\isacommand{by}\ (blast\ intro!:\ gcd_greatest\ intro:\ dvd_trans)
-\end{isabelle}
-This theorem concisely expresses the correctness of the \isa{gcd}
-function.
-We state it with the \isa{iff} attribute so that
-Isabelle can use it to remove some occurrences of \isa{gcd}.
-The theorem has a one-line
-proof using \isa{blast} supplied with two additional introduction
-rules. The exclamation mark
-({\isa{intro}}{\isa{!}})\ signifies safe rules, which are
-applied aggressively. Rules given without the exclamation mark
-are applied reluctantly and their uses can be undone if
-the search backtracks. Here the unsafe rule expresses transitivity
-of the divides relation:
-\begin{isabelle}
-\isasymlbrakk?m\ dvd\ ?n;\ ?n\ dvd\ ?p\isasymrbrakk\ \isasymLongrightarrow\ ?m\ dvd\ ?p%
-\rulename{dvd_trans}
-\end{isabelle}
-Applying \isa{dvd_trans} as
-an introduction rule entails a risk of looping, for it multiplies
-occurrences of the divides symbol. However, this proof relies
-on transitivity reasoning. The rule {\isa{gcd\_greatest}} is safe to apply
-aggressively because it yields simpler subgoals. The proof implicitly
-uses \isa{gcd_dvd1} and \isa{gcd_dvd2} as safe rules, because they were
-declared using \isa{iff}.%
-\index{Euclid's algorithm|)}\index{*gcd (constant)|)}\index{divides relation|)}
--- a/doc-src/TutorialI/Sets/sets.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1069 +0,0 @@
-\chapter{Sets, Functions and Relations}
-
-This chapter describes the formalization of typed set theory, which is
-the basis of much else in HOL\@. For example, an
-inductive definition yields a set, and the abstract theories of relations
-regard a relation as a set of pairs. The chapter introduces the well-known
-constants such as union and intersection, as well as the main operations on relations, such as converse, composition and
-transitive closure. Functions are also covered. They are not sets in
-HOL, but many of their properties concern sets: the range of a
-function is a set, and the inverse image of a function maps sets to sets.
-
-This chapter will be useful to anybody who plans to develop a substantial
-proof. Sets are convenient for formalizing computer science concepts such
-as grammars, logical calculi and state transition systems. Isabelle can
-prove many statements involving sets automatically.
-
-This chapter ends with a case study concerning model checking for the
-temporal logic CTL\@. Most of the other examples are simple. The
-chapter presents a small selection of built-in theorems in order to point
-out some key properties of the various constants and to introduce you to
-the notation.
-
-Natural deduction rules are provided for the set theory constants, but they
-are seldom used directly, so only a few are presented here.
-
-
-\section{Sets}
-
-\index{sets|(}%
-HOL's set theory should not be confused with traditional, untyped set
-theory, in which everything is a set. Our sets are typed. In a given set,
-all elements have the same type, say~$\tau$, and the set itself has type
-$\tau$~\tydx{set}.
-
-We begin with \textbf{intersection}, \textbf{union} and
-\textbf{complement}. In addition to the
-\textbf{membership relation}, there is a symbol for its negation. These
-points can be seen below.
-
-Here are the natural deduction rules for \rmindex{intersection}. Note
-the resemblance to those for conjunction.
-\begin{isabelle}
-\isasymlbrakk c\ \isasymin\ A;\ c\ \isasymin\ B\isasymrbrakk\
-\isasymLongrightarrow\ c\ \isasymin\ A\ \isasyminter\ B%
-\rulenamedx{IntI}\isanewline
-c\ \isasymin\ A\ \isasyminter\ B\ \isasymLongrightarrow\ c\ \isasymin\ A
-\rulenamedx{IntD1}\isanewline
-c\ \isasymin\ A\ \isasyminter\ B\ \isasymLongrightarrow\ c\ \isasymin\ B
-\rulenamedx{IntD2}
-\end{isabelle}
-
-Here are two of the many installed theorems concerning set
-complement.\index{complement!of a set}
-Note that it is denoted by a minus sign.
-\begin{isabelle}
-(c\ \isasymin\ -\ A)\ =\ (c\ \isasymnotin\ A)
-\rulenamedx{Compl_iff}\isanewline
--\ (A\ \isasymunion\ B)\ =\ -\ A\ \isasyminter\ -\ B
-\rulename{Compl_Un}
-\end{isabelle}
-
-Set \textbf{difference}\indexbold{difference!of sets} is the intersection
-of a set with the complement of another set. Here we also see the syntax
-for the empty set and for the universal set.
-\begin{isabelle}
-A\ \isasyminter\ (B\ -\ A)\ =\ \isacharbraceleft\isacharbraceright
-\rulename{Diff_disjoint}\isanewline
-A\ \isasymunion\ -\ A\ =\ UNIV%
-\rulename{Compl_partition}
-\end{isabelle}
-
-The \bfindex{subset relation} holds between two sets just if every element
-of one is also an element of the other. This relation is reflexive. These
-are its natural deduction rules:
-\begin{isabelle}
-({\isasymAnd}x.\ x\ \isasymin\ A\ \isasymLongrightarrow\ x\ \isasymin\ B)\ \isasymLongrightarrow\ A\ \isasymsubseteq\ B%
-\rulenamedx{subsetI}%
-\par\smallskip% \isanewline didn't leave enough space
-\isasymlbrakk A\ \isasymsubseteq\ B;\ c\ \isasymin\
-A\isasymrbrakk\ \isasymLongrightarrow\ c\
-\isasymin\ B%
-\rulenamedx{subsetD}
-\end{isabelle}
-In harder proofs, you may need to apply \isa{subsetD} giving a specific term
-for~\isa{c}. However, \isa{blast} can instantly prove facts such as this
-one:
-\begin{isabelle}
-(A\ \isasymunion\ B\ \isasymsubseteq\ C)\ =\
-(A\ \isasymsubseteq\ C\ \isasymand\ B\ \isasymsubseteq\ C)
-\rulenamedx{Un_subset_iff}
-\end{isabelle}
-Here is another example, also proved automatically:
-\begin{isabelle}
-\isacommand{lemma}\ "(A\
-\isasymsubseteq\ -B)\ =\ (B\ \isasymsubseteq\ -A)"\isanewline
-\isacommand{by}\ blast
-\end{isabelle}
-%
-This is the same example using \textsc{ascii} syntax, illustrating a pitfall:
-\begin{isabelle}
-\isacommand{lemma}\ "(A\ <=\ -B)\ =\ (B\ <=\ -A)"
-\end{isabelle}
-%
-The proof fails. It is not a statement about sets, due to overloading;
-the relation symbol~\isa{<=} can be any relation, not just
-subset.
-In this general form, the statement is not valid. Putting
-in a type constraint forces the variables to denote sets, allowing the
-proof to succeed:
-
-\begin{isabelle}
-\isacommand{lemma}\ "((A::\ {\isacharprime}a\ set)\ <=\ -B)\ =\ (B\ <=\
--A)"
-\end{isabelle}
-Section~\ref{sec:axclass} below describes overloading. Incidentally,
-\isa{A~\isasymsubseteq~-B} asserts that the sets \isa{A} and \isa{B} are
-disjoint.
-
-\medskip
-Two sets are \textbf{equal}\indexbold{equality!of sets} if they contain the
-same elements. This is
-the principle of \textbf{extensionality}\indexbold{extensionality!for sets}
-for sets.
-\begin{isabelle}
-({\isasymAnd}x.\ (x\ {\isasymin}\ A)\ =\ (x\ {\isasymin}\ B))\
-{\isasymLongrightarrow}\ A\ =\ B
-\rulenamedx{set_ext}
-\end{isabelle}
-Extensionality can be expressed as
-$A=B\iff (A\subseteq B)\conj (B\subseteq A)$.
-The following rules express both
-directions of this equivalence. Proving a set equation using
-\isa{equalityI} allows the two inclusions to be proved independently.
-\begin{isabelle}
-\isasymlbrakk A\ \isasymsubseteq\ B;\ B\ \isasymsubseteq\
-A\isasymrbrakk\ \isasymLongrightarrow\ A\ =\ B
-\rulenamedx{equalityI}
-\par\smallskip% \isanewline didn't leave enough space
-\isasymlbrakk A\ =\ B;\ \isasymlbrakk A\ \isasymsubseteq\ B;\ B\
-\isasymsubseteq\ A\isasymrbrakk\ \isasymLongrightarrow\ P\isasymrbrakk\
-\isasymLongrightarrow\ P%
-\rulenamedx{equalityE}
-\end{isabelle}
-
-
-\subsection{Finite Set Notation}
-
-\indexbold{sets!notation for finite}
-Finite sets are expressed using the constant \cdx{insert}, which is
-a form of union:
-\begin{isabelle}
-insert\ a\ A\ =\ \isacharbraceleft a\isacharbraceright\ \isasymunion\ A
-\rulename{insert_is_Un}
-\end{isabelle}
-%
-The finite set expression \isa{\isacharbraceleft
-a,b\isacharbraceright} abbreviates
-\isa{insert\ a\ (insert\ b\ \isacharbraceleft\isacharbraceright)}.
-Many facts about finite sets can be proved automatically:
-\begin{isabelle}
-\isacommand{lemma}\
-"\isacharbraceleft a,b\isacharbraceright\
-\isasymunion\ \isacharbraceleft c,d\isacharbraceright\ =\
-\isacharbraceleft a,b,c,d\isacharbraceright"\isanewline
-\isacommand{by}\ blast
-\end{isabelle}
-
-
-Not everything that we would like to prove is valid.
-Consider this attempt:
-\begin{isabelle}
-\isacommand{lemma}\ "\isacharbraceleft a,b\isacharbraceright\ \isasyminter\ \isacharbraceleft b,c\isacharbraceright\ =\
-\isacharbraceleft b\isacharbraceright"\isanewline
-\isacommand{apply}\ auto
-\end{isabelle}
-%
-The proof fails, leaving the subgoal \isa{b=c}. To see why it
-fails, consider a correct version:
-\begin{isabelle}
-\isacommand{lemma}\ "\isacharbraceleft a,b\isacharbraceright\ \isasyminter\
-\isacharbraceleft b,c\isacharbraceright\ =\ (if\ a=c\ then\
-\isacharbraceleft a,b\isacharbraceright\ else\ \isacharbraceleft
-b\isacharbraceright)"\isanewline
-\isacommand{apply}\ simp\isanewline
-\isacommand{by}\ blast
-\end{isabelle}
-
-Our mistake was to suppose that the various items were distinct. Another
-remark: this proof uses two methods, namely {\isa{simp}} and
-{\isa{blast}}. Calling {\isa{simp}} eliminates the
-\isa{if}-\isa{then}-\isa{else} expression, which {\isa{blast}}
-cannot break down. The combined methods (namely {\isa{force}} and
-\isa{auto}) can prove this fact in one step.
-
-
-\subsection{Set Comprehension}
-
-\index{set comprehensions|(}%
-The set comprehension \isa{\isacharbraceleft x.\
-P\isacharbraceright} expresses the set of all elements that satisfy the
-predicate~\isa{P}. Two laws describe the relationship between set
-comprehension and the membership relation:
-\begin{isabelle}
-(a\ \isasymin\
-\isacharbraceleft x.\ P\ x\isacharbraceright)\ =\ P\ a
-\rulename{mem_Collect_eq}\isanewline
-\isacharbraceleft x.\ x\ \isasymin\ A\isacharbraceright\ =\ A
-\rulename{Collect_mem_eq}
-\end{isabelle}
-
-\smallskip
-Facts such as these have trivial proofs:
-\begin{isabelle}
-\isacommand{lemma}\ "\isacharbraceleft x.\ P\ x\ \isasymor\
-x\ \isasymin\ A\isacharbraceright\ =\
-\isacharbraceleft x.\ P\ x\isacharbraceright\ \isasymunion\ A"
-\par\smallskip
-\isacommand{lemma}\ "\isacharbraceleft x.\ P\ x\
-\isasymlongrightarrow\ Q\ x\isacharbraceright\ =\
--\isacharbraceleft x.\ P\ x\isacharbraceright\
-\isasymunion\ \isacharbraceleft x.\ Q\ x\isacharbraceright"
-\end{isabelle}
-
-\smallskip
-Isabelle has a general syntax for comprehension, which is best
-described through an example:
-\begin{isabelle}
-\isacommand{lemma}\ "\isacharbraceleft p*q\ \isacharbar\ p\ q.\
-p{\isasymin}prime\ \isasymand\ q{\isasymin}prime\isacharbraceright\ =\
-\isanewline
-\ \ \ \ \ \ \ \ \isacharbraceleft z.\ \isasymexists p\ q.\ z\ =\ p*q\
-\isasymand\ p{\isasymin}prime\ \isasymand\
-q{\isasymin}prime\isacharbraceright"
-\end{isabelle}
-The left and right hand sides
-of this equation are identical. The syntax used in the
-left-hand side abbreviates the right-hand side: in this case, all numbers
-that are the product of two primes. The syntax provides a neat
-way of expressing any set given by an expression built up from variables
-under specific constraints. The drawback is that it hides the true form of
-the expression, with its existential quantifiers.
-
-\smallskip
-\emph{Remark}. We do not need sets at all. They are essentially equivalent
-to predicate variables, which are allowed in higher-order logic. The main
-benefit of sets is their notation; we can write \isa{x{\isasymin}A}
-and
-\isa{\isacharbraceleft z.\ P\isacharbraceright} where predicates would
-require writing
-\isa{A(x)} and
-\isa{{\isasymlambda}z.\ P}.
-\index{set comprehensions|)}
-
-
-\subsection{Binding Operators}
-
-\index{quantifiers!for sets|(}%
-Universal and existential quantifications may range over sets,
-with the obvious meaning. Here are the natural deduction rules for the
-bounded universal quantifier. Occasionally you will need to apply
-\isa{bspec} with an explicit instantiation of the variable~\isa{x}:
-%
-\begin{isabelle}
-({\isasymAnd}x.\ x\ \isasymin\ A\ \isasymLongrightarrow\ P\ x)\ \isasymLongrightarrow\ {\isasymforall}x\isasymin
-A.\ P\ x%
-\rulenamedx{ballI}%
-\isanewline
-\isasymlbrakk{\isasymforall}x\isasymin A.\
-P\ x;\ x\ \isasymin\
-A\isasymrbrakk\ \isasymLongrightarrow\ P\
-x%
-\rulenamedx{bspec}
-\end{isabelle}
-%
-Dually, here are the natural deduction rules for the
-bounded existential quantifier. You may need to apply
-\isa{bexI} with an explicit instantiation:
-\begin{isabelle}
-\isasymlbrakk P\ x;\
-x\ \isasymin\ A\isasymrbrakk\
-\isasymLongrightarrow\
-\isasymexists x\isasymin A.\ P\
-x%
-\rulenamedx{bexI}%
-\isanewline
-\isasymlbrakk\isasymexists x\isasymin A.\
-P\ x;\ {\isasymAnd}x.\
-{\isasymlbrakk}x\ \isasymin\ A;\
-P\ x\isasymrbrakk\ \isasymLongrightarrow\
-Q\isasymrbrakk\ \isasymLongrightarrow\ Q%
-\rulenamedx{bexE}
-\end{isabelle}
-\index{quantifiers!for sets|)}
-
-\index{union!indexed}%
-Unions can be formed over the values of a given set. The syntax is
-\mbox{\isa{\isasymUnion x\isasymin A.\ B}} or
-\isa{UN x:A.\ B} in \textsc{ascii}. Indexed union satisfies this basic law:
-\begin{isabelle}
-(b\ \isasymin\
-(\isasymUnion x\isasymin A. B\ x)) =\ (\isasymexists x\isasymin A.\
-b\ \isasymin\ B\ x)
-\rulenamedx{UN_iff}
-\end{isabelle}
-It has two natural deduction rules similar to those for the existential
-quantifier. Sometimes \isa{UN_I} must be applied explicitly:
-\begin{isabelle}
-\isasymlbrakk a\ \isasymin\
-A;\ b\ \isasymin\
-B\ a\isasymrbrakk\ \isasymLongrightarrow\
-b\ \isasymin\
-(\isasymUnion x\isasymin A. B\ x)
-\rulenamedx{UN_I}%
-\isanewline
-\isasymlbrakk b\ \isasymin\
-(\isasymUnion x\isasymin A. B\ x);\
-{\isasymAnd}x.\ {\isasymlbrakk}x\ \isasymin\
-A;\ b\ \isasymin\
-B\ x\isasymrbrakk\ \isasymLongrightarrow\
-R\isasymrbrakk\ \isasymLongrightarrow\ R%
-\rulenamedx{UN_E}
-\end{isabelle}
-%
-The following built-in abbreviation (see {\S}\ref{sec:abbreviations})
-lets us express the union over a \emph{type}:
-\begin{isabelle}
-\ \ \ \ \
-({\isasymUnion}x.\ B\ x)\ {\isasymequiv}\
-({\isasymUnion}x{\isasymin}UNIV. B\ x)
-\end{isabelle}
-
-We may also express the union of a set of sets, written \isa{Union\ C} in
-\textsc{ascii}:
-\begin{isabelle}
-(A\ \isasymin\ \isasymUnion C)\ =\ (\isasymexists X\isasymin C.\ A\
-\isasymin\ X)
-\rulenamedx{Union_iff}
-\end{isabelle}
-
-\index{intersection!indexed}%
-Intersections are treated dually, although they seem to be used less often
-than unions. The syntax below would be \isa{INT
-x:\ A.\ B} and \isa{Inter\ C} in \textsc{ascii}. Among others, these
-theorems are available:
-\begin{isabelle}
-(b\ \isasymin\
-(\isasymInter x\isasymin A. B\ x))\
-=\
-({\isasymforall}x\isasymin A.\
-b\ \isasymin\ B\ x)
-\rulenamedx{INT_iff}%
-\isanewline
-(A\ \isasymin\
-\isasymInter C)\ =\
-({\isasymforall}X\isasymin C.\
-A\ \isasymin\ X)
-\rulenamedx{Inter_iff}
-\end{isabelle}
-
-Isabelle uses logical equivalences such as those above in automatic proof.
-Unions, intersections and so forth are not simply replaced by their
-definitions. Instead, membership tests are simplified. For example, $x\in
-A\cup B$ is replaced by $x\in A\lor x\in B$.
-
-The internal form of a comprehension involves the constant
-\cdx{Collect},
-which occasionally appears when a goal or theorem
-is displayed. For example, \isa{Collect\ P} is the same term as
-\isa{\isacharbraceleft x.\ P\ x\isacharbraceright}. The same thing can
-happen with quantifiers: \hbox{\isa{All\ P}}\index{*All (constant)}
-is
-\isa{{\isasymforall}x.\ P\ x} and
-\hbox{\isa{Ex\ P}}\index{*Ex (constant)} is \isa{\isasymexists x.\ P\ x};
-also \isa{Ball\ A\ P}\index{*Ball (constant)} is
-\isa{{\isasymforall}x\isasymin A.\ P\ x} and
-\isa{Bex\ A\ P}\index{*Bex (constant)} is
-\isa{\isasymexists x\isasymin A.\ P\ x}. For indexed unions and
-intersections, you may see the constants
-\cdx{UNION} and \cdx{INTER}\@.
-The internal constant for $\varepsilon x.P(x)$ is~\cdx{Eps}.
-
-We have only scratched the surface of Isabelle/HOL's set theory, which provides
-hundreds of theorems for your use.
-
-
-\subsection{Finiteness and Cardinality}
-
-\index{sets!finite}%
-The predicate \sdx{finite} holds of all finite sets. Isabelle/HOL
-includes many familiar theorems about finiteness and cardinality
-(\cdx{card}). For example, we have theorems concerning the
-cardinalities of unions, intersections and the
-powerset:\index{cardinality}
-%
-\begin{isabelle}
-{\isasymlbrakk}finite\ A;\ finite\ B\isasymrbrakk\isanewline
-\isasymLongrightarrow\ card\ A\ \isacharplus\ card\ B\ =\ card\ (A\ \isasymunion\ B)\ \isacharplus\ card\ (A\ \isasyminter\ B)
-\rulenamedx{card_Un_Int}%
-\isanewline
-\isanewline
-finite\ A\ \isasymLongrightarrow\ card\
-(Pow\ A)\ =\ 2\ \isacharcircum\ card\ A%
-\rulenamedx{card_Pow}%
-\isanewline
-\isanewline
-finite\ A\ \isasymLongrightarrow\isanewline
-card\ \isacharbraceleft B.\ B\ \isasymsubseteq\
-A\ \isasymand\ card\ B\ =\
-k\isacharbraceright\ =\ card\ A\ choose\ k%
-\rulename{n_subsets}
-\end{isabelle}
-Writing $|A|$ as $n$, the last of these theorems says that the number of
-$k$-element subsets of~$A$ is \index{binomial coefficients}
-$\binom{n}{k}$.
-
-%\begin{warn}
-%The term \isa{finite\ A} is defined via a syntax translation as an
-%abbreviation for \isa{A {\isasymin} Finites}, where the constant
-%\cdx{Finites} denotes the set of all finite sets of a given type. There
-%is no constant \isa{finite}.
-%\end{warn}
-\index{sets|)}
-
-
-\section{Functions}
-
-\index{functions|(}%
-This section describes a few concepts that involve
-functions. Some of the more important theorems are given along with the
-names. A few sample proofs appear. Unlike with set theory, however,
-we cannot simply state lemmas and expect them to be proved using
-\isa{blast}.
-
-\subsection{Function Basics}
-
-Two functions are \textbf{equal}\indexbold{equality!of functions}
-if they yield equal results given equal
-arguments. This is the principle of
-\textbf{extensionality}\indexbold{extensionality!for functions} for
-functions:
-\begin{isabelle}
-({\isasymAnd}x.\ f\ x\ =\ g\ x)\ {\isasymLongrightarrow}\ f\ =\ g
-\rulenamedx{ext}
-\end{isabelle}
-
-\indexbold{updating a function}%
-Function \textbf{update} is useful for modelling machine states. It has
-the obvious definition and many useful facts are proved about
-it. In particular, the following equation is installed as a simplification
-rule:
-\begin{isabelle}
-(f(x:=y))\ z\ =\ (if\ z\ =\ x\ then\ y\ else\ f\ z)
-\rulename{fun_upd_apply}
-\end{isabelle}
-Two syntactic points must be noted. In
-\isa{(f(x:=y))\ z} we are applying an updated function to an
-argument; the outer parentheses are essential. A series of two or more
-updates can be abbreviated as shown on the left-hand side of this theorem:
-\begin{isabelle}
-f(x:=y,\ x:=z)\ =\ f(x:=z)
-\rulename{fun_upd_upd}
-\end{isabelle}
-Note also that we can write \isa{f(x:=z)} with only one pair of parentheses
-when it is not being applied to an argument.
-
-\medskip
-The \bfindex{identity function} and function
-\textbf{composition}\indexbold{composition!of functions} are
-defined:
-\begin{isabelle}%
-id\ \isasymequiv\ {\isasymlambda}x.\ x%
-\rulenamedx{id_def}\isanewline
-f\ \isasymcirc\ g\ \isasymequiv\
-{\isasymlambda}x.\ f\
-(g\ x)%
-\rulenamedx{o_def}
-\end{isabelle}
-%
-Many familiar theorems concerning the identity and composition
-are proved. For example, we have the associativity of composition:
-\begin{isabelle}
-f\ \isasymcirc\ (g\ \isasymcirc\ h)\ =\ f\ \isasymcirc\ g\ \isasymcirc\ h
-\rulename{o_assoc}
-\end{isabelle}
-
-\subsection{Injections, Surjections, Bijections}
-
-\index{injections}\index{surjections}\index{bijections}%
-A function may be \textbf{injective}, \textbf{surjective} or \textbf{bijective}:
-\begin{isabelle}
-inj_on\ f\ A\ \isasymequiv\ {\isasymforall}x\isasymin A.\
-{\isasymforall}y\isasymin A.\ f\ x\ =\ f\ y\ \isasymlongrightarrow\ x\
-=\ y%
-\rulenamedx{inj_on_def}\isanewline
-surj\ f\ \isasymequiv\ {\isasymforall}y.\
-\isasymexists x.\ y\ =\ f\ x%
-\rulenamedx{surj_def}\isanewline
-bij\ f\ \isasymequiv\ inj\ f\ \isasymand\ surj\ f
-\rulenamedx{bij_def}
-\end{isabelle}
-The second argument
-of \isa{inj_on} lets us express that a function is injective over a
-given set. This refinement is useful in higher-order logic, where
-functions are total; in some cases, a function's natural domain is a subset
-of its domain type. Writing \isa{inj\ f} abbreviates \isa{inj_on\ f\
-UNIV}, for when \isa{f} is injective everywhere.
-
-The operator \isa{inv} expresses the
-\textbf{inverse}\indexbold{inverse!of a function}
-of a function. In
-general the inverse may not be well behaved. We have the usual laws,
-such as these:
-\begin{isabelle}
-inj\ f\ \ \isasymLongrightarrow\ inv\ f\ (f\ x)\ =\ x%
-\rulename{inv_f_f}\isanewline
-surj\ f\ \isasymLongrightarrow\ f\ (inv\ f\ y)\ =\ y
-\rulename{surj_f_inv_f}\isanewline
-bij\ f\ \ \isasymLongrightarrow\ inv\ (inv\ f)\ =\ f
-\rulename{inv_inv_eq}
-\end{isabelle}
-%
-%Other useful facts are that the inverse of an injection
-%is a surjection and vice versa; the inverse of a bijection is
-%a bijection.
-%\begin{isabelle}
-%inj\ f\ \isasymLongrightarrow\ surj\
-%(inv\ f)
-%\rulename{inj_imp_surj_inv}\isanewline
-%surj\ f\ \isasymLongrightarrow\ inj\ (inv\ f)
-%\rulename{surj_imp_inj_inv}\isanewline
-%bij\ f\ \isasymLongrightarrow\ bij\ (inv\ f)
-%\rulename{bij_imp_bij_inv}
-%\end{isabelle}
-%
-%The converses of these results fail. Unless a function is
-%well behaved, little can be said about its inverse. Here is another
-%law:
-%\begin{isabelle}
-%{\isasymlbrakk}bij\ f;\ bij\ g\isasymrbrakk\ \isasymLongrightarrow\ inv\ (f\ \isasymcirc\ g)\ =\ inv\ g\ \isasymcirc\ inv\ f%
-%\rulename{o_inv_distrib}
-%\end{isabelle}
-
-Theorems involving these concepts can be hard to prove. The following
-example is easy, but it cannot be proved automatically. To begin
-with, we need a law that relates the equality of functions to
-equality over all arguments:
-\begin{isabelle}
-(f\ =\ g)\ =\ ({\isasymforall}x.\ f\ x\ =\ g\ x)
-\rulename{fun_eq_iff}
-\end{isabelle}
-%
-This is just a restatement of
-extensionality.\indexbold{extensionality!for functions}
-Our lemma
-states that an injection can be cancelled from the left side of
-function composition:
-\begin{isabelle}
-\isacommand{lemma}\ "inj\ f\ \isasymLongrightarrow\ (f\ o\ g\ =\ f\ o\ h)\ =\ (g\ =\ h)"\isanewline
-\isacommand{apply}\ (simp\ add:\ fun_eq_iff\ inj_on_def)\isanewline
-\isacommand{apply}\ auto\isanewline
-\isacommand{done}
-\end{isabelle}
-
-The first step of the proof invokes extensionality and the definitions
-of injectiveness and composition. It leaves one subgoal:
-\begin{isabelle}
-\ 1.\ {\isasymforall}x\ y.\ f\ x\ =\ f\ y\ \isasymlongrightarrow\ x\ =\ y\
-\isasymLongrightarrow\isanewline
-\ \ \ \ ({\isasymforall}x.\ f\ (g\ x)\ =\ f\ (h\ x))\ =\ ({\isasymforall}x.\ g\ x\ =\ h\ x)
-\end{isabelle}
-This can be proved using the \isa{auto} method.
-
-
-\subsection{Function Image}
-
-The \textbf{image}\indexbold{image!under a function}
-of a set under a function is a most useful notion. It
-has the obvious definition:
-\begin{isabelle}
-f\ `\ A\ \isasymequiv\ \isacharbraceleft y.\ \isasymexists x\isasymin
-A.\ y\ =\ f\ x\isacharbraceright
-\rulenamedx{image_def}
-\end{isabelle}
-%
-Here are some of the many facts proved about image:
-\begin{isabelle}
-(f\ \isasymcirc\ g)\ `\ r\ =\ f\ `\ g\ `\ r
-\rulename{image_compose}\isanewline
-f`(A\ \isasymunion\ B)\ =\ f`A\ \isasymunion\ f`B
-\rulename{image_Un}\isanewline
-inj\ f\ \isasymLongrightarrow\ f`(A\ \isasyminter\
-B)\ =\ f`A\ \isasyminter\ f`B
-\rulename{image_Int}
-%\isanewline
-%bij\ f\ \isasymLongrightarrow\ f\ `\ (-\ A)\ =\ -\ f\ `\ A%
-%\rulename{bij_image_Compl_eq}
-\end{isabelle}
-
-
-Laws involving image can often be proved automatically. Here
-are two examples, illustrating connections with indexed union and with the
-general syntax for comprehension:
-\begin{isabelle}
-\isacommand{lemma}\ "f`A\ \isasymunion\ g`A\ =\ ({\isasymUnion}x{\isasymin}A.\ \isacharbraceleft f\ x,\ g\
-x\isacharbraceright)"
-\par\smallskip
-\isacommand{lemma}\ "f\ `\ \isacharbraceleft(x,y){.}\ P\ x\ y\isacharbraceright\ =\ \isacharbraceleft f(x,y)\ \isacharbar\ x\ y.\ P\ x\
-y\isacharbraceright"
-\end{isabelle}
-
-\medskip
-\index{range!of a function}%
-A function's \textbf{range} is the set of values that the function can
-take on. It is, in fact, the image of the universal set under
-that function. There is no constant \isa{range}. Instead,
-\sdx{range} abbreviates an application of image to \isa{UNIV}:
-\begin{isabelle}
-\ \ \ \ \ range\ f\
-{\isasymrightleftharpoons}\ f`UNIV
-\end{isabelle}
-%
-Few theorems are proved specifically
-for {\isa{range}}; in most cases, you should look for a more general
-theorem concerning images.
-
-\medskip
-\textbf{Inverse image}\index{inverse image!of a function} is also useful.
-It is defined as follows:
-\begin{isabelle}
-f\ -`\ B\ \isasymequiv\ \isacharbraceleft x.\ f\ x\ \isasymin\ B\isacharbraceright
-\rulenamedx{vimage_def}
-\end{isabelle}
-%
-This is one of the facts proved about it:
-\begin{isabelle}
-f\ -`\ (-\ A)\ =\ -\ f\ -`\ A%
-\rulename{vimage_Compl}
-\end{isabelle}
-\index{functions|)}
-
-
-\section{Relations}
-\label{sec:Relations}
-
-\index{relations|(}%
-A \textbf{relation} is a set of pairs. As such, the set operations apply
-to them. For instance, we may form the union of two relations. Other
-primitives are defined specifically for relations.
-
-\subsection{Relation Basics}
-
-The \bfindex{identity relation}, also known as equality, has the obvious
-definition:
-\begin{isabelle}
-Id\ \isasymequiv\ \isacharbraceleft p.\ \isasymexists x.\ p\ =\ (x,x){\isacharbraceright}%
-\rulenamedx{Id_def}
-\end{isabelle}
-
-\indexbold{composition!of relations}%
-\textbf{Composition} of relations (the infix \sdx{O}) is also
-available:
-\begin{isabelle}
-r\ O\ s\ \isasymequiv\ \isacharbraceleft(x,z).\ \isasymexists y.\ (x,y)\ \isasymin\ s\ \isasymand\ (y,z)\ \isasymin\ r\isacharbraceright
-\rulenamedx{rel_comp_def}
-\end{isabelle}
-%
-This is one of the many lemmas proved about these concepts:
-\begin{isabelle}
-R\ O\ Id\ =\ R
-\rulename{R_O_Id}
-\end{isabelle}
-%
-Composition is monotonic, as are most of the primitives appearing
-in this chapter. We have many theorems similar to the following
-one:
-\begin{isabelle}
-\isasymlbrakk r\isacharprime\ \isasymsubseteq\ r;\ s\isacharprime\
-\isasymsubseteq\ s\isasymrbrakk\ \isasymLongrightarrow\ r\isacharprime\ O\
-s\isacharprime\ \isasymsubseteq\ r\ O\ s%
-\rulename{rel_comp_mono}
-\end{isabelle}
-
-\indexbold{converse!of a relation}%
-\indexbold{inverse!of a relation}%
-The \textbf{converse} or inverse of a
-relation exchanges the roles of the two operands. We use the postfix
-notation~\isa{r\isasyminverse} or
-\isa{r\isacharcircum-1} in ASCII\@.
-\begin{isabelle}
-((a,b)\ \isasymin\ r\isasyminverse)\ =\
-((b,a)\ \isasymin\ r)
-\rulenamedx{converse_iff}
-\end{isabelle}
-%
-Here is a typical law proved about converse and composition:
-\begin{isabelle}
-(r\ O\ s)\isasyminverse\ =\ s\isasyminverse\ O\ r\isasyminverse
-\rulename{converse_rel_comp}
-\end{isabelle}
-
-\indexbold{image!under a relation}%
-The \textbf{image} of a set under a relation is defined
-analogously to image under a function:
-\begin{isabelle}
-(b\ \isasymin\ r\ ``\ A)\ =\ (\isasymexists x\isasymin
-A.\ (x,b)\ \isasymin\ r)
-\rulenamedx{Image_iff}
-\end{isabelle}
-It satisfies many similar laws.
-
-\index{domain!of a relation}%
-\index{range!of a relation}%
-The \textbf{domain} and \textbf{range} of a relation are defined in the
-standard way:
-\begin{isabelle}
-(a\ \isasymin\ Domain\ r)\ =\ (\isasymexists y.\ (a,y)\ \isasymin\
-r)
-\rulenamedx{Domain_iff}%
-\isanewline
-(a\ \isasymin\ Range\ r)\
-\ =\ (\isasymexists y.\
-(y,a)\
-\isasymin\ r)
-\rulenamedx{Range_iff}
-\end{isabelle}
-
-Iterated composition of a relation is available. The notation overloads
-that of exponentiation. Two simplification rules are installed:
-\begin{isabelle}
-R\ \isacharcircum\ \isadigit{0}\ =\ Id\isanewline
-R\ \isacharcircum\ Suc\ n\ =\ R\ O\ R\isacharcircum n
-\end{isabelle}
-
-\subsection{The Reflexive and Transitive Closure}
-
-\index{reflexive and transitive closure|(}%
-The \textbf{reflexive and transitive closure} of the
-relation~\isa{r} is written with a
-postfix syntax. In ASCII we write \isa{r\isacharcircum*} and in
-symbol notation~\isa{r\isactrlsup *}. It is the least solution of the
-equation
-\begin{isabelle}
-r\isactrlsup *\ =\ Id\ \isasymunion \ (r\ O\ r\isactrlsup *)
-\rulename{rtrancl_unfold}
-\end{isabelle}
-%
-Among its basic properties are three that serve as introduction
-rules:
-\begin{isabelle}
-(a,\ a)\ \isasymin \ r\isactrlsup *
-\rulenamedx{rtrancl_refl}\isanewline
-p\ \isasymin \ r\ \isasymLongrightarrow \ p\ \isasymin \ r\isactrlsup *
-\rulenamedx{r_into_rtrancl}\isanewline
-\isasymlbrakk (a,b)\ \isasymin \ r\isactrlsup *;\
-(b,c)\ \isasymin \ r\isactrlsup *\isasymrbrakk \ \isasymLongrightarrow \
-(a,c)\ \isasymin \ r\isactrlsup *
-\rulenamedx{rtrancl_trans}
-\end{isabelle}
-%
-Induction over the reflexive transitive closure is available:
-\begin{isabelle}
-\isasymlbrakk (a,\ b)\ \isasymin \ r\isactrlsup *;\ P\ a;\ \isasymAnd y\ z.\ \isasymlbrakk (a,\ y)\ \isasymin \ r\isactrlsup *;\ (y,\ z)\ \isasymin \ r;\ P\ y\isasymrbrakk \ \isasymLongrightarrow \ P\ z\isasymrbrakk \isanewline
-\isasymLongrightarrow \ P\ b%
-\rulename{rtrancl_induct}
-\end{isabelle}
-%
-Idempotence is one of the laws proved about the reflexive transitive
-closure:
-\begin{isabelle}
-(r\isactrlsup *)\isactrlsup *\ =\ r\isactrlsup *
-\rulename{rtrancl_idemp}
-\end{isabelle}
-
-\smallskip
-The transitive closure is similar. The ASCII syntax is
-\isa{r\isacharcircum+}. It has two introduction rules:
-\begin{isabelle}
-p\ \isasymin \ r\ \isasymLongrightarrow \ p\ \isasymin \ r\isactrlsup +
-\rulenamedx{r_into_trancl}\isanewline
-\isasymlbrakk (a,\ b)\ \isasymin \ r\isactrlsup +;\ (b,\ c)\ \isasymin \ r\isactrlsup +\isasymrbrakk \ \isasymLongrightarrow \ (a,\ c)\ \isasymin \ r\isactrlsup +
-\rulenamedx{trancl_trans}
-\end{isabelle}
-%
-The induction rule resembles the one shown above.
-A typical lemma states that transitive closure commutes with the converse
-operator:
-\begin{isabelle}
-(r\isasyminverse )\isactrlsup +\ =\ (r\isactrlsup +)\isasyminverse
-\rulename{trancl_converse}
-\end{isabelle}
-
-\subsection{A Sample Proof}
-
-The reflexive transitive closure also commutes with the converse
-operator. Let us examine the proof. Each direction of the equivalence
-is proved separately. The two proofs are almost identical. Here
-is the first one:
-\begin{isabelle}
-\isacommand{lemma}\ rtrancl_converseD:\ "(x,y)\ \isasymin \
-(r\isasyminverse)\isactrlsup *\ \isasymLongrightarrow \ (y,x)\ \isasymin
-\ r\isactrlsup *"\isanewline
-\isacommand{apply}\ (erule\ rtrancl_induct)\isanewline
-\ \isacommand{apply}\ (rule\ rtrancl_refl)\isanewline
-\isacommand{apply}\ (blast\ intro:\ rtrancl_trans)\isanewline
-\isacommand{done}
-\end{isabelle}
-%
-The first step of the proof applies induction, leaving these subgoals:
-\begin{isabelle}
-\ 1.\ (x,\ x)\ \isasymin \ r\isactrlsup *\isanewline
-\ 2.\ \isasymAnd y\ z.\ \isasymlbrakk (x,y)\ \isasymin \
-(r\isasyminverse)\isactrlsup *;\ (y,z)\ \isasymin \ r\isasyminverse ;\
-(y,x)\ \isasymin \ r\isactrlsup *\isasymrbrakk \isanewline
-\ \ \ \ \ \ \ \ \ \ \isasymLongrightarrow \ (z,x)\ \isasymin \ r\isactrlsup *
-\end{isabelle}
-%
-The first subgoal is trivial by reflexivity. The second follows
-by first eliminating the converse operator, yielding the
-assumption \isa{(z,y)\
-\isasymin\ r}, and then
-applying the introduction rules shown above. The same proof script handles
-the other direction:
-\begin{isabelle}
-\isacommand{lemma}\ rtrancl_converseI:\ "(y,x)\ \isasymin \ r\isactrlsup *\ \isasymLongrightarrow \ (x,y)\ \isasymin \ (r\isasyminverse)\isactrlsup *"\isanewline
-\isacommand{apply}\ (erule\ rtrancl_induct)\isanewline
-\ \isacommand{apply}\ (rule\ rtrancl_refl)\isanewline
-\isacommand{apply}\ (blast\ intro:\ rtrancl_trans)\isanewline
-\isacommand{done}
-\end{isabelle}
-
-
-Finally, we combine the two lemmas to prove the desired equation:
-\begin{isabelle}
-\isacommand{lemma}\ rtrancl_converse:\ "(r\isasyminverse)\isactrlsup *\ =\ (r\isactrlsup *)\isasyminverse"\isanewline
-\isacommand{by}\ (auto\ intro:\ rtrancl_converseI\ dest:\
-rtrancl_converseD)
-\end{isabelle}
-
-\begin{warn}
-This trivial proof requires \isa{auto} rather than \isa{blast} because
-of a subtle issue involving ordered pairs. Here is a subgoal that
-arises internally after the rules
-\isa{equalityI} and \isa{subsetI} have been applied:
-\begin{isabelle}
-\ 1.\ \isasymAnd x.\ x\ \isasymin \ (r\isasyminverse )\isactrlsup *\ \isasymLongrightarrow \ x\ \isasymin \ (r\isactrlsup
-*)\isasyminverse
-%ignore subgoal 2
-%\ 2.\ \isasymAnd x.\ x\ \isasymin \ (r\isactrlsup *)\isasyminverse \
-%\isasymLongrightarrow \ x\ \isasymin \ (r\isasyminverse )\isactrlsup *
-\end{isabelle}
-\par\noindent
-We cannot apply \isa{rtrancl_converseD}\@. It refers to
-ordered pairs, while \isa{x} is a variable of product type.
-The \isa{simp} and \isa{blast} methods can do nothing, so let us try
-\isa{clarify}:
-\begin{isabelle}
-\ 1.\ \isasymAnd a\ b.\ (a,b)\ \isasymin \ (r\isasyminverse )\isactrlsup *\ \isasymLongrightarrow \ (b,a)\ \isasymin \ r\isactrlsup
-*
-\end{isabelle}
-Now that \isa{x} has been replaced by the pair \isa{(a,b)}, we can
-proceed. Other methods that split variables in this way are \isa{force},
-\isa{auto}, \isa{fast} and \isa{best}. Section~\ref{sec:products} will discuss proof
-techniques for ordered pairs in more detail.
-\end{warn}
-\index{relations|)}\index{reflexive and transitive closure|)}
-
-
-\section{Well-Founded Relations and Induction}
-\label{sec:Well-founded}
-
-\index{relations!well-founded|(}%
-A well-founded relation captures the notion of a terminating
-process. Complex recursive functions definitions must specify a
-well-founded relation that justifies their
-termination~\cite{isabelle-function}. Most of the forms of induction found
-in mathematics are merely special cases of induction over a
-well-founded relation.
-
-Intuitively, the relation~$\prec$ is \textbf{well-founded} if it admits no
-infinite descending chains
-\[ \cdots \prec a@2 \prec a@1 \prec a@0. \]
-Well-foundedness can be hard to show. The various
-formulations are all complicated. However, often a relation
-is well-founded by construction. HOL provides
-theorems concerning ways of constructing a well-founded relation. The
-most familiar way is to specify a
-\index{measure functions}\textbf{measure function}~\isa{f} into
-the natural numbers, when $\isa{x}\prec \isa{y}\iff \isa{f x} < \isa{f y}$;
-we write this particular relation as
-\isa{measure~f}.
-
-\begin{warn}
-You may want to skip the rest of this section until you need to perform a
-complex recursive function definition or induction. The induction rule
-returned by
-\isacommand{fun} is good enough for most purposes. We use an explicit
-well-founded induction only in {\S}\ref{sec:CTL-revisited}.
-\end{warn}
-
-Isabelle/HOL declares \cdx{less_than} as a relation object,
-that is, a set of pairs of natural numbers. Two theorems tell us that this
-relation behaves as expected and that it is well-founded:
-\begin{isabelle}
-((x,y)\ \isasymin\ less_than)\ =\ (x\ <\ y)
-\rulenamedx{less_than_iff}\isanewline
-wf\ less_than
-\rulenamedx{wf_less_than}
-\end{isabelle}
-
-The notion of measure generalizes to the
-\index{inverse image!of a relation}\textbf{inverse image} of
-a relation. Given a relation~\isa{r} and a function~\isa{f}, we express a
-new relation using \isa{f} as a measure. An infinite descending chain on
-this new relation would give rise to an infinite descending chain
-on~\isa{r}. Isabelle/HOL defines this concept and proves a
-theorem stating that it preserves well-foundedness:
-\begin{isabelle}
-inv_image\ r\ f\ \isasymequiv\ \isacharbraceleft(x,y).\ (f\ x,\ f\ y)\
-\isasymin\ r\isacharbraceright
-\rulenamedx{inv_image_def}\isanewline
-wf\ r\ \isasymLongrightarrow\ wf\ (inv_image\ r\ f)
-\rulenamedx{wf_inv_image}
-\end{isabelle}
-
-A measure function involves the natural numbers. The relation \isa{measure
-size} justifies primitive recursion and structural induction over a
-datatype. Isabelle/HOL defines
-\isa{measure} as shown:
-\begin{isabelle}
-measure\ \isasymequiv\ inv_image\ less_than%
-\rulenamedx{measure_def}\isanewline
-wf\ (measure\ f)
-\rulenamedx{wf_measure}
-\end{isabelle}
-
-Of the other constructions, the most important is the
-\bfindex{lexicographic product} of two relations. It expresses the
-standard dictionary ordering over pairs. We write \isa{ra\ <*lex*>\
-rb}, where \isa{ra} and \isa{rb} are the two operands. The
-lexicographic product satisfies the usual definition and it preserves
-well-foundedness:
-\begin{isabelle}
-ra\ <*lex*>\ rb\ \isasymequiv \isanewline
-\ \ \isacharbraceleft ((a,b),(a',b')).\ (a,a')\ \isasymin \ ra\
-\isasymor\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \,a=a'\ \isasymand \ (b,b')\
-\isasymin \ rb\isacharbraceright
-\rulenamedx{lex_prod_def}%
-\par\smallskip
-\isasymlbrakk wf\ ra;\ wf\ rb\isasymrbrakk \ \isasymLongrightarrow \ wf\ (ra\ <*lex*>\ rb)
-\rulenamedx{wf_lex_prod}
-\end{isabelle}
-
-%These constructions can be used in a
-%\textbf{recdef} declaration ({\S}\ref{sec:recdef-simplification}) to define
-%the well-founded relation used to prove termination.
-
-The \bfindex{multiset ordering}, useful for hard termination proofs, is
-available in the Library~\cite{HOL-Library}.
-Baader and Nipkow \cite[{\S}2.5]{Baader-Nipkow} discuss it.
-
-\medskip
-Induction\index{induction!well-founded|(}
-comes in many forms,
-including traditional mathematical induction, structural induction on
-lists and induction on size. All are instances of the following rule,
-for a suitable well-founded relation~$\prec$:
-\[ \infer{P(a)}{\infer*{P(x)}{[\forall y.\, y\prec x \imp P(y)]}} \]
-To show $P(a)$ for a particular term~$a$, it suffices to show $P(x)$ for
-arbitrary~$x$ under the assumption that $P(y)$ holds for $y\prec x$.
-Intuitively, the well-foundedness of $\prec$ ensures that the chains of
-reasoning are finite.
-
-\smallskip
-In Isabelle, the induction rule is expressed like this:
-\begin{isabelle}
-{\isasymlbrakk}wf\ r;\
- {\isasymAnd}x.\ {\isasymforall}y.\ (y,x)\ \isasymin\ r\
-\isasymlongrightarrow\ P\ y\ \isasymLongrightarrow\ P\ x\isasymrbrakk\
-\isasymLongrightarrow\ P\ a
-\rulenamedx{wf_induct}
-\end{isabelle}
-Here \isa{wf\ r} expresses that the relation~\isa{r} is well-founded.
-
-Many familiar induction principles are instances of this rule.
-For example, the predecessor relation on the natural numbers
-is well-founded; induction over it is mathematical induction.
-The ``tail of'' relation on lists is well-founded; induction over
-it is structural induction.%
-\index{induction!well-founded|)}%
-\index{relations!well-founded|)}
-
-
-\section{Fixed Point Operators}
-
-\index{fixed points|(}%
-Fixed point operators define sets
-recursively. They are invoked implicitly when making an inductive
-definition, as discussed in Chap.\ts\ref{chap:inductive} below. However,
-they can be used directly, too. The
-\emph{least} or \emph{strongest} fixed point yields an inductive
-definition; the \emph{greatest} or \emph{weakest} fixed point yields a
-coinductive definition. Mathematicians may wish to note that the
-existence of these fixed points is guaranteed by the Knaster-Tarski
-theorem.
-
-\begin{warn}
-Casual readers should skip the rest of this section. We use fixed point
-operators only in {\S}\ref{sec:VMC}.
-\end{warn}
-
-The theory applies only to monotonic functions.\index{monotone functions|bold}
-Isabelle's definition of monotone is overloaded over all orderings:
-\begin{isabelle}
-mono\ f\ \isasymequiv\ {\isasymforall}A\ B.\ A\ \isasymle\ B\ \isasymlongrightarrow\ f\ A\ \isasymle\ f\ B%
-\rulenamedx{mono_def}
-\end{isabelle}
-%
-For fixed point operators, the ordering will be the subset relation: if
-$A\subseteq B$ then we expect $f(A)\subseteq f(B)$. In addition to its
-definition, monotonicity has the obvious introduction and destruction
-rules:
-\begin{isabelle}
-({\isasymAnd}A\ B.\ A\ \isasymle\ B\ \isasymLongrightarrow\ f\ A\ \isasymle\ f\ B)\ \isasymLongrightarrow\ mono\ f%
-\rulename{monoI}%
-\par\smallskip% \isanewline didn't leave enough space
-{\isasymlbrakk}mono\ f;\ A\ \isasymle\ B\isasymrbrakk\
-\isasymLongrightarrow\ f\ A\ \isasymle\ f\ B%
-\rulename{monoD}
-\end{isabelle}
-
-The most important properties of the least fixed point are that
-it is a fixed point and that it enjoys an induction rule:
-\begin{isabelle}
-mono\ f\ \isasymLongrightarrow\ lfp\ f\ =\ f\ (lfp\ f)
-\rulename{lfp_unfold}%
-\par\smallskip% \isanewline didn't leave enough space
-{\isasymlbrakk}a\ \isasymin\ lfp\ f;\ mono\ f;\isanewline
- \ {\isasymAnd}x.\ x\
-\isasymin\ f\ (lfp\ f\ \isasyminter\ \isacharbraceleft x.\ P\
-x\isacharbraceright)\ \isasymLongrightarrow\ P\ x\isasymrbrakk\
-\isasymLongrightarrow\ P\ a%
-\rulename{lfp_induct}
-\end{isabelle}
-%
-The induction rule shown above is more convenient than the basic
-one derived from the minimality of {\isa{lfp}}. Observe that both theorems
-demand \isa{mono\ f} as a premise.
-
-The greatest fixed point is similar, but it has a \bfindex{coinduction} rule:
-\begin{isabelle}
-mono\ f\ \isasymLongrightarrow\ gfp\ f\ =\ f\ (gfp\ f)
-\rulename{gfp_unfold}%
-\isanewline
-{\isasymlbrakk}mono\ f;\ a\ \isasymin\ X;\ X\ \isasymsubseteq\ f\ (X\
-\isasymunion\ gfp\ f)\isasymrbrakk\ \isasymLongrightarrow\ a\ \isasymin\
-gfp\ f%
-\rulename{coinduct}
-\end{isabelle}
-A \textbf{bisimulation}\index{bisimulations}
-is perhaps the best-known concept defined as a
-greatest fixed point. Exhibiting a bisimulation to prove the equality of
-two agents in a process algebra is an example of coinduction.
-The coinduction rule can be strengthened in various ways.
-\index{fixed points|)}
-
-%The section "Case Study: Verified Model Checking" is part of this chapter
-\input{CTL/ctl}
-\endinput
--- a/doc-src/TutorialI/ToyList/ToyList.thy Tue Aug 28 13:15:15 2012 +0200
+++ b/doc-src/TutorialI/ToyList/ToyList.thy Tue Aug 28 14:37:57 2012 +0200
@@ -2,6 +2,17 @@
imports Datatype
begin
+(*<*)
+ML {*
+ let
+ val texts =
+ map (File.read o Path.append (Thy_Load.master_directory @{theory}) o Path.explode)
+ ["ToyList1", "ToyList2"];
+ val trs = Outer_Syntax.parse Position.start (implode texts);
+ in @{assert} (Toplevel.is_toplevel (fold Toplevel.command trs Toplevel.toplevel)) end;
+*}
+(*>*)
+
text{*\noindent
HOL already has a predefined theory of lists called @{text List} ---
@{text ToyList} is merely a small fragment of it chosen as an example. In
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/TutorialI/ToyList/ToyList1 Tue Aug 28 14:37:57 2012 +0200
@@ -0,0 +1,16 @@
+theory ToyList
+imports Datatype
+begin
+
+datatype 'a list = Nil ("[]")
+ | Cons 'a "'a list" (infixr "#" 65)
+
+(* This is the append function: *)
+primrec app :: "'a list => 'a list => 'a list" (infixr "@" 65)
+where
+"[] @ ys = ys" |
+"(x # xs) @ ys = x # (xs @ ys)"
+
+primrec rev :: "'a list => 'a list" where
+"rev [] = []" |
+"rev (x # xs) = (rev xs) @ (x # [])"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/TutorialI/ToyList/ToyList2 Tue Aug 28 14:37:57 2012 +0200
@@ -0,0 +1,21 @@
+lemma app_Nil2 [simp]: "xs @ [] = xs"
+apply(induct_tac xs)
+apply(auto)
+done
+
+lemma app_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
+apply(induct_tac xs)
+apply(auto)
+done
+
+lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)"
+apply(induct_tac xs)
+apply(auto)
+done
+
+theorem rev_rev [simp]: "rev(rev xs) = xs"
+apply(induct_tac xs)
+apply(auto)
+done
+
+end
--- a/doc-src/TutorialI/ToyList2/ToyList.thy Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,37 +0,0 @@
-theory ToyList
-imports Datatype
-begin
-
-datatype 'a list = Nil ("[]")
- | Cons 'a "'a list" (infixr "#" 65)
-
-(* This is the append function: *)
-primrec app :: "'a list => 'a list => 'a list" (infixr "@" 65)
-where
-"[] @ ys = ys" |
-"(x # xs) @ ys = x # (xs @ ys)"
-
-primrec rev :: "'a list => 'a list" where
-"rev [] = []" |
-"rev (x # xs) = (rev xs) @ (x # [])"
-lemma app_Nil2 [simp]: "xs @ [] = xs"
-apply(induct_tac xs)
-apply(auto)
-done
-
-lemma app_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
-apply(induct_tac xs)
-apply(auto)
-done
-
-lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)"
-apply(induct_tac xs)
-apply(auto)
-done
-
-theorem rev_rev [simp]: "rev(rev xs) = xs"
-apply(induct_tac xs)
-apply(auto)
-done
-
-end
--- a/doc-src/TutorialI/ToyList2/ToyList1 Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,16 +0,0 @@
-theory ToyList
-imports Datatype
-begin
-
-datatype 'a list = Nil ("[]")
- | Cons 'a "'a list" (infixr "#" 65)
-
-(* This is the append function: *)
-primrec app :: "'a list => 'a list => 'a list" (infixr "@" 65)
-where
-"[] @ ys = ys" |
-"(x # xs) @ ys = x # (xs @ ys)"
-
-primrec rev :: "'a list => 'a list" where
-"rev [] = []" |
-"rev (x # xs) = (rev xs) @ (x # [])"
--- a/doc-src/TutorialI/ToyList2/ToyList2 Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,21 +0,0 @@
-lemma app_Nil2 [simp]: "xs @ [] = xs"
-apply(induct_tac xs)
-apply(auto)
-done
-
-lemma app_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)"
-apply(induct_tac xs)
-apply(auto)
-done
-
-lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)"
-apply(induct_tac xs)
-apply(auto)
-done
-
-theorem rev_rev [simp]: "rev(rev xs) = xs"
-apply(induct_tac xs)
-apply(auto)
-done
-
-end
--- a/doc-src/TutorialI/Types/numerics.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,543 +0,0 @@
-\section{Numbers}
-\label{sec:numbers}
-
-\index{numbers|(}%
-Until now, our numerical examples have used the type of \textbf{natural
-numbers},
-\isa{nat}. This is a recursive datatype generated by the constructors
-zero and successor, so it works well with inductive proofs and primitive
-recursive function definitions. HOL also provides the type
-\isa{int} of \textbf{integers}, which lack induction but support true
-subtraction. With subtraction, arithmetic reasoning is easier, which makes
-the integers preferable to the natural numbers for
-complicated arithmetic expressions, even if they are non-negative. There are also the types
-\isa{rat}, \isa{real} and \isa{complex}: the rational, real and complex numbers. Isabelle has no
-subtyping, so the numeric
-types are distinct and there are functions to convert between them.
-Most numeric operations are overloaded: the same symbol can be
-used at all numeric types. Table~\ref{tab:overloading} in the appendix
-shows the most important operations, together with the priorities of the
-infix symbols. Algebraic properties are organized using type classes
-around algebraic concepts such as rings and fields;
-a property such as the commutativity of addition is a single theorem
-(\isa{add_commute}) that applies to all numeric types.
-
-\index{linear arithmetic}%
-Many theorems involving numeric types can be proved automatically by
-Isabelle's arithmetic decision procedure, the method
-\methdx{arith}. Linear arithmetic comprises addition, subtraction
-and multiplication by constant factors; subterms involving other operators
-are regarded as variables. The procedure can be slow, especially if the
-subgoal to be proved involves subtraction over type \isa{nat}, which
-causes case splits. On types \isa{nat} and \isa{int}, \methdx{arith}
-can deal with quantifiers---this is known as Presburger arithmetic---whereas on type \isa{real} it cannot.
-
-The simplifier reduces arithmetic expressions in other
-ways, such as dividing through by common factors. For problems that lie
-outside the scope of automation, HOL provides hundreds of
-theorems about multiplication, division, etc., that can be brought to
-bear. You can locate them using Proof General's Find
-button. A few lemmas are given below to show what
-is available.
-
-\subsection{Numeric Literals}
-\label{sec:numerals}
-
-\index{numeric literals|(}%
-The constants \cdx{0} and \cdx{1} are overloaded. They denote zero and one,
-respectively, for all numeric types. Other values are expressed by numeric
-literals, which consist of one or more decimal digits optionally preceeded by a minus sign (\isa{-}). Examples are \isa{2}, \isa{-3} and
-\isa{441223334678}. Literals are available for the types of natural
-numbers, integers, rationals, reals, etc.; they denote integer values of
-arbitrary size.
-
-Literals look like constants, but they abbreviate
-terms representing the number in a two's complement binary notation.
-Isabelle performs arithmetic on literals by rewriting rather
-than using the hardware arithmetic. In most cases arithmetic
-is fast enough, even for numbers in the millions. The arithmetic operations
-provided for literals include addition, subtraction, multiplication,
-integer division and remainder. Fractions of literals (expressed using
-division) are reduced to lowest terms.
-
-\begin{warn}\index{overloading!and arithmetic}
-The arithmetic operators are
-overloaded, so you must be careful to ensure that each numeric
-expression refers to a specific type, if necessary by inserting
-type constraints. Here is an example of what can go wrong:
-\par
-\begin{isabelle}
-\isacommand{lemma}\ "2\ *\ m\ =\ m\ +\ m"
-\end{isabelle}
-%
-Carefully observe how Isabelle displays the subgoal:
-\begin{isabelle}
-\ 1.\ (2::'a)\ *\ m\ =\ m\ +\ m
-\end{isabelle}
-The type \isa{'a} given for the literal \isa{2} warns us that no numeric
-type has been specified. The problem is underspecified. Given a type
-constraint such as \isa{nat}, \isa{int} or \isa{real}, it becomes trivial.
-\end{warn}
-
-\begin{warn}
-\index{function@\isacommand {function} (command)!and numeric literals}
-Numeric literals are not constructors and therefore
-must not be used in patterns. For example, this declaration is
-rejected:
-\begin{isabelle}
-\isacommand{function}\ h\ \isakeyword{where}\isanewline
-"h\ 3\ =\ 2"\isanewline
-\isacharbar "h\ i\ \ =\ i"
-\end{isabelle}
-
-You should use a conditional expression instead:
-\begin{isabelle}
-"h\ i\ =\ (if\ i\ =\ 3\ then\ 2\ else\ i)"
-\end{isabelle}
-\index{numeric literals|)}
-\end{warn}
-
-
-\subsection{The Type of Natural Numbers, {\tt\slshape nat}}
-
-\index{natural numbers|(}\index{*nat (type)|(}%
-This type requires no introduction: we have been using it from the
-beginning. Hundreds of theorems about the natural numbers are
-proved in the theories \isa{Nat} and \isa{Divides}.
-Basic properties of addition and multiplication are available through the
-axiomatic type class for semirings (\S\ref{sec:numeric-classes}).
-
-\subsubsection{Literals}
-\index{numeric literals!for type \protect\isa{nat}}%
-The notational options for the natural numbers are confusing. Recall that an
-overloaded constant can be defined independently for each type; the definition
-of \cdx{1} for type \isa{nat} is
-\begin{isabelle}
-1\ \isasymequiv\ Suc\ 0
-\rulename{One_nat_def}
-\end{isabelle}
-This is installed as a simplification rule, so the simplifier will replace
-every occurrence of \isa{1::nat} by \isa{Suc\ 0}. Literals are obviously
-better than nested \isa{Suc}s at expressing large values. But many theorems,
-including the rewrite rules for primitive recursive functions, can only be
-applied to terms of the form \isa{Suc\ $n$}.
-
-The following default simplification rules replace
-small literals by zero and successor:
-\begin{isabelle}
-2\ +\ n\ =\ Suc\ (Suc\ n)
-\rulename{add_2_eq_Suc}\isanewline
-n\ +\ 2\ =\ Suc\ (Suc\ n)
-\rulename{add_2_eq_Suc'}
-\end{isabelle}
-It is less easy to transform \isa{100} into \isa{Suc\ 99} (for example), and
-the simplifier will normally reverse this transformation. Novices should
-express natural numbers using \isa{0} and \isa{Suc} only.
-
-\subsubsection{Division}
-\index{division!for type \protect\isa{nat}}%
-The infix operators \isa{div} and \isa{mod} are overloaded.
-Isabelle/HOL provides the basic facts about quotient and remainder
-on the natural numbers:
-\begin{isabelle}
-m\ mod\ n\ =\ (if\ m\ <\ n\ then\ m\ else\ (m\ -\ n)\ mod\ n)
-\rulename{mod_if}\isanewline
-m\ div\ n\ *\ n\ +\ m\ mod\ n\ =\ m%
-\rulenamedx{mod_div_equality}
-\end{isabelle}
-
-Many less obvious facts about quotient and remainder are also provided.
-Here is a selection:
-\begin{isabelle}
-a\ *\ b\ div\ c\ =\ a\ *\ (b\ div\ c)\ +\ a\ *\ (b\ mod\ c)\ div\ c%
-\rulename{div_mult1_eq}\isanewline
-a\ *\ b\ mod\ c\ =\ a\ *\ (b\ mod\ c)\ mod\ c%
-\rulename{mod_mult_right_eq}\isanewline
-a\ div\ (b*c)\ =\ a\ div\ b\ div\ c%
-\rulename{div_mult2_eq}\isanewline
-a\ mod\ (b*c)\ =\ b * (a\ div\ b\ mod\ c)\ +\ a\ mod\ b%
-\rulename{mod_mult2_eq}\isanewline
-0\ <\ c\ \isasymLongrightarrow \ (c\ *\ a)\ div\ (c\ *\ b)\ =\ a\ div\ b%
-\rulename{div_mult_mult1}\isanewline
-(m\ mod\ n)\ *\ k\ =\ (m\ *\ k)\ mod\ (n\ *\ k)
-\rulenamedx{mod_mult_distrib}\isanewline
-m\ \isasymle \ n\ \isasymLongrightarrow \ m\ div\ k\ \isasymle \ n\ div\ k%
-\rulename{div_le_mono}
-\end{isabelle}
-
-Surprisingly few of these results depend upon the
-divisors' being nonzero.
-\index{division!by zero}%
-That is because division by
-zero yields zero:
-\begin{isabelle}
-a\ div\ 0\ =\ 0
-\rulename{DIVISION_BY_ZERO_DIV}\isanewline
-a\ mod\ 0\ =\ a%
-\rulename{DIVISION_BY_ZERO_MOD}
-\end{isabelle}
-In \isa{div_mult_mult1} above, one of
-the two divisors (namely~\isa{c}) must still be nonzero.
-
-The \textbf{divides} relation\index{divides relation}
-has the standard definition, which
-is overloaded over all numeric types:
-\begin{isabelle}
-m\ dvd\ n\ \isasymequiv\ {\isasymexists}k.\ n\ =\ m\ *\ k
-\rulenamedx{dvd_def}
-\end{isabelle}
-%
-Section~\ref{sec:proving-euclid} discusses proofs involving this
-relation. Here are some of the facts proved about it:
-\begin{isabelle}
-\isasymlbrakk m\ dvd\ n;\ n\ dvd\ m\isasymrbrakk \ \isasymLongrightarrow \ m\ =\ n%
-\rulenamedx{dvd_antisym}\isanewline
-\isasymlbrakk k\ dvd\ m;\ k\ dvd\ n\isasymrbrakk \ \isasymLongrightarrow \ k\ dvd\ (m\ +\ n)
-\rulenamedx{dvd_add}
-\end{isabelle}
-
-\subsubsection{Subtraction}
-
-There are no negative natural numbers, so \isa{m\ -\ n} equals zero unless
-\isa{m} exceeds~\isa{n}. The following is one of the few facts
-about \isa{m\ -\ n} that is not subject to
-the condition \isa{n\ \isasymle \ m}.
-\begin{isabelle}
-(m\ -\ n)\ *\ k\ =\ m\ *\ k\ -\ n\ *\ k%
-\rulenamedx{diff_mult_distrib}
-\end{isabelle}
-Natural number subtraction has few
-nice properties; often you should remove it by simplifying with this split
-rule.
-\begin{isabelle}
-P(a-b)\ =\ ((a<b\ \isasymlongrightarrow \ P\
-0)\ \isasymand \ (\isasymforall d.\ a\ =\ b+d\ \isasymlongrightarrow \ P\
-d))
-\rulename{nat_diff_split}
-\end{isabelle}
-For example, splitting helps to prove the following fact.
-\begin{isabelle}
-\isacommand{lemma}\ "(n\ -\ 2)\ *\ (n\ +\ 2)\ =\ n\ *\ n\ -\ (4::nat)"\isanewline
-\isacommand{apply}\ (simp\ split:\ nat_diff_split,\ clarify)\isanewline
-\ 1.\ \isasymAnd d.\ \isasymlbrakk n\ <\ 2;\ n\ *\ n\ =\ 4\ +\ d\isasymrbrakk \ \isasymLongrightarrow \ d\ =\ 0
-\end{isabelle}
-The result lies outside the scope of linear arithmetic, but
- it is easily found
-if we explicitly split \isa{n<2} as \isa{n=0} or \isa{n=1}:
-\begin{isabelle}
-\isacommand{apply}\ (subgoal_tac\ "n=0\ |\ n=1",\ force,\ arith)\isanewline
-\isacommand{done}
-\end{isabelle}%%%%%%
-\index{natural numbers|)}\index{*nat (type)|)}
-
-
-\subsection{The Type of Integers, {\tt\slshape int}}
-
-\index{integers|(}\index{*int (type)|(}%
-Reasoning methods for the integers resemble those for the natural numbers,
-but induction and
-the constant \isa{Suc} are not available. HOL provides many lemmas for
-proving inequalities involving integer multiplication and division, similar
-to those shown above for type~\isa{nat}. The laws of addition, subtraction
-and multiplication are available through the axiomatic type class for rings
-(\S\ref{sec:numeric-classes}).
-
-The \rmindex{absolute value} function \cdx{abs} is overloaded, and is
-defined for all types that involve negative numbers, including the integers.
-The \isa{arith} method can prove facts about \isa{abs} automatically,
-though as it does so by case analysis, the cost can be exponential.
-\begin{isabelle}
-\isacommand{lemma}\ "abs\ (x+y)\ \isasymle \ abs\ x\ +\ abs\ (y\ ::\ int)"\isanewline
-\isacommand{by}\ arith
-\end{isabelle}
-
-For division and remainder,\index{division!by negative numbers}
-the treatment of negative divisors follows
-mathematical practice: the sign of the remainder follows that
-of the divisor:
-\begin{isabelle}
-0\ <\ b\ \isasymLongrightarrow \ 0\ \isasymle \ a\ mod\ b%
-\rulename{pos_mod_sign}\isanewline
-0\ <\ b\ \isasymLongrightarrow \ a\ mod\ b\ <\ b%
-\rulename{pos_mod_bound}\isanewline
-b\ <\ 0\ \isasymLongrightarrow \ a\ mod\ b\ \isasymle \ 0
-\rulename{neg_mod_sign}\isanewline
-b\ <\ 0\ \isasymLongrightarrow \ b\ <\ a\ mod\ b%
-\rulename{neg_mod_bound}
-\end{isabelle}
-ML treats negative divisors in the same way, but most computer hardware
-treats signed operands using the same rules as for multiplication.
-Many facts about quotients and remainders are provided:
-\begin{isabelle}
-(a\ +\ b)\ div\ c\ =\isanewline
-a\ div\ c\ +\ b\ div\ c\ +\ (a\ mod\ c\ +\ b\ mod\ c)\ div\ c%
-\rulename{zdiv_zadd1_eq}
-\par\smallskip
-(a\ +\ b)\ mod\ c\ =\ (a\ mod\ c\ +\ b\ mod\ c)\ mod\ c%
-\rulename{mod_add_eq}
-\end{isabelle}
-
-\begin{isabelle}
-(a\ *\ b)\ div\ c\ =\ a\ *\ (b\ div\ c)\ +\ a\ *\ (b\ mod\ c)\ div\ c%
-\rulename{zdiv_zmult1_eq}\isanewline
-(a\ *\ b)\ mod\ c\ =\ a\ *\ (b\ mod\ c)\ mod\ c%
-\rulename{zmod_zmult1_eq}
-\end{isabelle}
-
-\begin{isabelle}
-0\ <\ c\ \isasymLongrightarrow \ a\ div\ (b*c)\ =\ a\ div\ b\ div\ c%
-\rulename{zdiv_zmult2_eq}\isanewline
-0\ <\ c\ \isasymLongrightarrow \ a\ mod\ (b*c)\ =\ b*(a\ div\ b\ mod\
-c)\ +\ a\ mod\ b%
-\rulename{zmod_zmult2_eq}
-\end{isabelle}
-The last two differ from their natural number analogues by requiring
-\isa{c} to be positive. Since division by zero yields zero, we could allow
-\isa{c} to be zero. However, \isa{c} cannot be negative: a counterexample
-is
-$\isa{a} = 7$, $\isa{b} = 2$ and $\isa{c} = -3$, when the left-hand side of
-\isa{zdiv_zmult2_eq} is $-2$ while the right-hand side is~$-1$.
-The prefix~\isa{z} in many theorem names recalls the use of $\mathbb{Z}$ to
-denote the set of integers.%
-\index{integers|)}\index{*int (type)|)}
-
-Induction is less important for integers than it is for the natural numbers, but it can be valuable if the range of integers has a lower or upper bound. There are four rules for integer induction, corresponding to the possible relations of the bound ($\geq$, $>$, $\leq$ and $<$):
-\begin{isabelle}
-\isasymlbrakk k\ \isasymle \ i;\ P\ k;\ \isasymAnd i.\ \isasymlbrakk k\ \isasymle \ i;\ P\ i\isasymrbrakk \ \isasymLongrightarrow \ P(i+1)\isasymrbrakk \ \isasymLongrightarrow \ P\ i%
-\rulename{int_ge_induct}\isanewline
-\isasymlbrakk k\ <\ i;\ P(k+1);\ \isasymAnd i.\ \isasymlbrakk k\ <\ i;\ P\ i\isasymrbrakk \ \isasymLongrightarrow \ P(i+1)\isasymrbrakk \ \isasymLongrightarrow \ P\ i%
-\rulename{int_gr_induct}\isanewline
-\isasymlbrakk i\ \isasymle \ k;\ P\ k;\ \isasymAnd i.\ \isasymlbrakk i\ \isasymle \ k;\ P\ i\isasymrbrakk \ \isasymLongrightarrow \ P(i-1)\isasymrbrakk \ \isasymLongrightarrow \ P\ i%
-\rulename{int_le_induct}\isanewline
-\isasymlbrakk i\ <\ k;\ P(k-1);\ \isasymAnd i.\ \isasymlbrakk i\ <\ k;\ P\ i\isasymrbrakk \ \isasymLongrightarrow \ P(i-1)\isasymrbrakk \ \isasymLongrightarrow \ P\ i%
-\rulename{int_less_induct}
-\end{isabelle}
-
-
-\subsection{The Types of Rational, Real and Complex Numbers}
-\label{sec:real}
-
-\index{rational numbers|(}\index{*rat (type)|(}%
-\index{real numbers|(}\index{*real (type)|(}%
-\index{complex numbers|(}\index{*complex (type)|(}%
-These types provide true division, the overloaded operator \isa{/},
-which differs from the operator \isa{div} of the
-natural numbers and integers. The rationals and reals are
-\textbf{dense}: between every two distinct numbers lies another.
-This property follows from the division laws, since if $x\not=y$ then $(x+y)/2$ lies between them:
-\begin{isabelle}
-a\ <\ b\ \isasymLongrightarrow \ \isasymexists r.\ a\ <\ r\ \isasymand \ r\ <\ b%
-\rulename{dense}
-\end{isabelle}
-
-The real numbers are, moreover, \textbf{complete}: every set of reals that
-is bounded above has a least upper bound. Completeness distinguishes the
-reals from the rationals, for which the set $\{x\mid x^2<2\}$ has no least
-upper bound. (It could only be $\surd2$, which is irrational.) The
-formalization of completeness, which is complicated,
-can be found in theory \texttt{RComplete}.
-
-Numeric literals\index{numeric literals!for type \protect\isa{real}}
-for type \isa{real} have the same syntax as those for type
-\isa{int} and only express integral values. Fractions expressed
-using the division operator are automatically simplified to lowest terms:
-\begin{isabelle}
-\ 1.\ P\ ((3\ /\ 4)\ *\ (8\ /\ 15))\isanewline
-\isacommand{apply} simp\isanewline
-\ 1.\ P\ (2\ /\ 5)
-\end{isabelle}
-Exponentiation can express floating-point values such as
-\isa{2 * 10\isacharcircum6}, which will be simplified to integers.
-
-\begin{warn}
-Types \isa{rat}, \isa{real} and \isa{complex} are provided by theory HOL-Complex, which is
-Main extended with a definitional development of the rational, real and complex
-numbers. Base your theory upon theory \thydx{Complex_Main}, not the
-usual \isa{Main}.%
-\end{warn}
-
-Available in the logic HOL-NSA is the
-theory \isa{Hyperreal}, which define the type \tydx{hypreal} of
-\rmindex{non-standard reals}. These
-\textbf{hyperreals} include infinitesimals, which represent infinitely
-small and infinitely large quantities; they facilitate proofs
-about limits, differentiation and integration~\cite{fleuriot-jcm}. The
-development defines an infinitely large number, \isa{omega} and an
-infinitely small positive number, \isa{epsilon}. The
-relation $x\approx y$ means ``$x$ is infinitely close to~$y$.''
-Theory \isa{Hyperreal} also defines transcendental functions such as sine,
-cosine, exponential and logarithm --- even the versions for type
-\isa{real}, because they are defined using nonstandard limits.%
-\index{rational numbers|)}\index{*rat (type)|)}%
-\index{real numbers|)}\index{*real (type)|)}%
-\index{complex numbers|)}\index{*complex (type)|)}
-
-
-\subsection{The Numeric Type Classes}\label{sec:numeric-classes}
-
-Isabelle/HOL organises its numeric theories using axiomatic type classes.
-Hundreds of basic properties are proved in the theory \isa{Ring_and_Field}.
-These lemmas are available (as simprules if they were declared as such)
-for all numeric types satisfying the necessary axioms. The theory defines
-dozens of type classes, such as the following:
-\begin{itemize}
-\item
-\tcdx{semiring} and \tcdx{ordered_semiring}: a \emph{semiring}
-provides the associative operators \isa{+} and~\isa{*}, with \isa{0} and~\isa{1}
-as their respective identities. The operators enjoy the usual distributive law,
-and addition (\isa{+}) is also commutative.
-An \emph{ordered semiring} is also linearly
-ordered, with addition and multiplication respecting the ordering. Type \isa{nat} is an ordered semiring.
-\item
-\tcdx{ring} and \tcdx{ordered_ring}: a \emph{ring} extends a semiring
-with unary minus (the additive inverse) and subtraction (both
-denoted~\isa{-}). An \emph{ordered ring} includes the absolute value
-function, \cdx{abs}. Type \isa{int} is an ordered ring.
-\item
-\tcdx{field} and \tcdx{ordered_field}: a field extends a ring with the
-multiplicative inverse (called simply \cdx{inverse} and division~(\isa{/})).
-An ordered field is based on an ordered ring. Type \isa{complex} is a field, while type \isa{real} is an ordered field.
-\item
-\tcdx{division_by_zero} includes all types where \isa{inverse 0 = 0}
-and \isa{a / 0 = 0}. These include all of Isabelle's standard numeric types.
-However, the basic properties of fields are derived without assuming
-division by zero.
-\end{itemize}
-
-Hundreds of basic lemmas are proved, each of which
-holds for all types in the corresponding type class. In most
-cases, it is obvious whether a property is valid for a particular type. No
-abstract properties involving subtraction hold for type \isa{nat};
-instead, theorems such as
-\isa{diff_mult_distrib} are proved specifically about subtraction on
-type~\isa{nat}. All abstract properties involving division require a field.
-Obviously, all properties involving orderings required an ordered
-structure.
-
-The class \tcdx{ring_no_zero_divisors} of rings without zero divisors satisfies a number of natural cancellation laws, the first of which characterizes this class:
-\begin{isabelle}
-(a\ *\ b\ =\ (0::'a))\ =\ (a\ =\ (0::'a)\ \isasymor \ b\ =\ (0::'a))
-\rulename{mult_eq_0_iff}\isanewline
-(a\ *\ c\ =\ b\ *\ c)\ =\ (c\ =\ (0::'a)\ \isasymor \ a\ =\ b)
-\rulename{mult_cancel_right}
-\end{isabelle}
-\begin{pgnote}
-Setting the flag \pgmenu{Isabelle} $>$ \pgmenu{Settings} $>$
-\pgmenu{Show Sorts} will display the type classes of all type variables.
-\end{pgnote}
-\noindent
-Here is how the theorem \isa{mult_cancel_left} appears with the flag set.
-\begin{isabelle}
-((c::'a::ring_no_zero_divisors)\ *\ (a::'a::ring_no_zero_divisors) =\isanewline
-\ c\ *\ (b::'a::ring_no_zero_divisors))\ =\isanewline
-(c\ =\ (0::'a::ring_no_zero_divisors)\ \isasymor\ a\ =\ b)
-\end{isabelle}
-
-
-\subsubsection{Simplifying with the AC-Laws}
-Suppose that two expressions are equal, differing only in
-associativity and commutativity of addition. Simplifying with the
-following equations sorts the terms and groups them to the right, making
-the two expressions identical.
-\begin{isabelle}
-a\ +\ b\ +\ c\ =\ a\ +\ (b\ +\ c)
-\rulenamedx{add_assoc}\isanewline
-a\ +\ b\ =\ b\ +\ a%
-\rulenamedx{add_commute}\isanewline
-a\ +\ (b\ +\ c)\ =\ b\ +\ (a\ +\ c)
-\rulename{add_left_commute}
-\end{isabelle}
-The name \isa{add_ac}\index{*add_ac (theorems)}
-refers to the list of all three theorems; similarly
-there is \isa{mult_ac}.\index{*mult_ac (theorems)}
-They are all proved for semirings and therefore hold for all numeric types.
-
-Here is an example of the sorting effect. Start
-with this goal, which involves type \isa{nat}.
-\begin{isabelle}
-\ 1.\ Suc\ (i\ +\ j\ *\ l\ *\ k\ +\ m\ *\ n)\ =\
-f\ (n\ *\ m\ +\ i\ +\ k\ *\ j\ *\ l)
-\end{isabelle}
-%
-Simplify using \isa{add_ac} and \isa{mult_ac}.
-\begin{isabelle}
-\isacommand{apply}\ (simp\ add:\ add_ac\ mult_ac)
-\end{isabelle}
-%
-Here is the resulting subgoal.
-\begin{isabelle}
-\ 1.\ Suc\ (i\ +\ (m\ *\ n\ +\ j\ *\ (k\ *\ l)))\
-=\ f\ (i\ +\ (m\ *\ n\ +\ j\ *\ (k\ *\ l)))%
-\end{isabelle}
-
-
-\subsubsection{Division Laws for Fields}
-
-Here is a selection of rules about the division operator. The following
-are installed as default simplification rules in order to express
-combinations of products and quotients as rational expressions:
-\begin{isabelle}
-a\ *\ (b\ /\ c)\ =\ a\ *\ b\ /\ c
-\rulename{times_divide_eq_right}\isanewline
-b\ /\ c\ *\ a\ =\ b\ *\ a\ /\ c
-\rulename{times_divide_eq_left}\isanewline
-a\ /\ (b\ /\ c)\ =\ a\ *\ c\ /\ b
-\rulename{divide_divide_eq_right}\isanewline
-a\ /\ b\ /\ c\ =\ a\ /\ (b\ *\ c)
-\rulename{divide_divide_eq_left}
-\end{isabelle}
-
-Signs are extracted from quotients in the hope that complementary terms can
-then be cancelled:
-\begin{isabelle}
--\ (a\ /\ b)\ =\ -\ a\ /\ b
-\rulename{minus_divide_left}\isanewline
--\ (a\ /\ b)\ =\ a\ /\ -\ b
-\rulename{minus_divide_right}
-\end{isabelle}
-
-The following distributive law is available, but it is not installed as a
-simplification rule.
-\begin{isabelle}
-(a\ +\ b)\ /\ c\ =\ a\ /\ c\ +\ b\ /\ c%
-\rulename{add_divide_distrib}
-\end{isabelle}
-
-
-\subsubsection{Absolute Value}
-
-The \rmindex{absolute value} function \cdx{abs} is available for all
-ordered rings, including types \isa{int}, \isa{rat} and \isa{real}.
-It satisfies many properties,
-such as the following:
-\begin{isabelle}
-\isasymbar x\ *\ y\isasymbar \ =\ \isasymbar x\isasymbar \ *\ \isasymbar y\isasymbar
-\rulename{abs_mult}\isanewline
-(\isasymbar a\isasymbar \ \isasymle \ b)\ =\ (a\ \isasymle \ b\ \isasymand \ -\ a\ \isasymle \ b)
-\rulename{abs_le_iff}\isanewline
-\isasymbar a\ +\ b\isasymbar \ \isasymle \ \isasymbar a\isasymbar \ +\ \isasymbar b\isasymbar
-\rulename{abs_triangle_ineq}
-\end{isabelle}
-
-\begin{warn}
-The absolute value bars shown above cannot be typed on a keyboard. They
-can be entered using the X-symbol package. In \textsc{ascii}, type \isa{abs x} to
-get \isa{\isasymbar x\isasymbar}.
-\end{warn}
-
-
-\subsubsection{Raising to a Power}
-
-Another type class, \tcdx{ordered\_idom}, specifies rings that also have
-exponentation to a natural number power, defined using the obvious primitive
-recursion. Theory \thydx{Power} proves various theorems, such as the
-following.
-\begin{isabelle}
-a\ \isacharcircum \ (m\ +\ n)\ =\ a\ \isacharcircum \ m\ *\ a\ \isacharcircum \ n%
-\rulename{power_add}\isanewline
-a\ \isacharcircum \ (m\ *\ n)\ =\ (a\ \isacharcircum \ m)\ \isacharcircum \ n%
-\rulename{power_mult}\isanewline
-\isasymbar a\ \isacharcircum \ n\isasymbar \ =\ \isasymbar a\isasymbar \ \isacharcircum \ n%
-\rulename{power_abs}
-\end{isabelle}%%%%%%%%%%%%%%%%%%%%%%%%%
-\index{numbers|)}
--- a/doc-src/TutorialI/Types/types.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,69 +0,0 @@
-\chapter{More about Types}
-\label{ch:more-types}
-
-So far we have learned about a few basic types (for example \isa{bool} and
-\isa{nat}), type abbreviations (\isacommand{types}) and recursive datatypes
-(\isacommand{datatype}). This chapter will introduce more
-advanced material:
-\begin{itemize}
-\item Pairs ({\S}\ref{sec:products}) and records ({\S}\ref{sec:records}),
-and how to reason about them.
-\item Type classes: how to specify and reason about axiomatic collections of
- types ({\S}\ref{sec:axclass}). This section leads on to a discussion of
- Isabelle's numeric types ({\S}\ref{sec:numbers}).
-\item Introducing your own types: how to define types that
- cannot be constructed with any of the basic methods
- ({\S}\ref{sec:adv-typedef}).
-\end{itemize}
-
-The material in this section goes beyond the needs of most novices.
-Serious users should at least skim the sections as far as type classes.
-That material is fairly advanced; read the beginning to understand what it
-is about, but consult the rest only when necessary.
-
-\index{pairs and tuples|(}
-\input{document/Pairs} %%%Section "Pairs and Tuples"
-\index{pairs and tuples|)}
-
-\input{document/Records} %%%Section "Records"
-
-
-\section{Type Classes} %%%Section
-\label{sec:axclass}
-\index{axiomatic type classes|(}
-\index{*axclass|(}
-
-The programming language Haskell has popularized the notion of type
-classes: a type class is a set of types with a
-common interface: all types in that class must provide the functions
-in the interface. Isabelle offers a similar type class concept: in
-addition, properties (\emph{class axioms}) can be specified which any
-instance of this type class must obey. Thus we can talk about a type
-$\tau$ being in a class $C$, which is written $\tau :: C$. This is the case
-if $\tau$ satisfies the axioms of $C$. Furthermore, type classes can be
-organized in a hierarchy. Thus there is the notion of a class $D$
-being a subclass\index{subclasses} of a class $C$, written $D
-< C$. This is the case if all axioms of $C$ are also provable in $D$.
-
-In this section we introduce the most important concepts behind type
-classes by means of a running example from algebra. This should give
-you an intuition how to use type classes and to understand
-specifications involving type classes. Type classes are covered more
-deeply in a separate tutorial \cite{isabelle-classes}.
-
-\subsection{Overloading}
-\label{sec:overloading}
-\index{overloading|(}
-
-\input{document/Overloading}
-
-\index{overloading|)}
-
-\input{document/Axioms}
-
-\index{type classes|)}
-\index{*class|)}
-
-\input{Types/numerics} %%%Section "Numbers"
-
-\input{document/Typedefs} %%%Section "Introducing New Types"
--- a/doc-src/TutorialI/appendix.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,190 +0,0 @@
-\appendix
-
-\chapter{Appendix}
-\label{sec:Appendix}
-
-\begin{table}[htbp]
-\begin{center}
-\begin{tabular}{|l|l|l|}
-\hline
-\indexboldpos{\isasymlbrakk}{$Isabrl} &
-\texttt{[|}\index{$Isabrl@\ttlbr|bold} &
-\verb$\<lbrakk>$ \\
-\indexboldpos{\isasymrbrakk}{$Isabrr} &
-\texttt{|]}\index{$Isabrr@\ttrbr|bold} &
-\verb$\<rbrakk>$ \\
-\indexboldpos{\isasymImp}{$IsaImp} &
-\ttindexboldpos{==>}{$IsaImp} &
-\verb$\<Longrightarrow>$ \\
-\isasymAnd\index{$IsaAnd@\isasymAnd|bold}&
-\texttt{!!}\index{$IsaAnd@\ttAnd|bold} &
-\verb$\<And>$ \\
-\indexboldpos{\isasymequiv}{$IsaEq} &
-\ttindexboldpos{==}{$IsaEq} &
-\verb$\<equiv>$ \\
-\indexboldpos{\isasymrightleftharpoons}{$IsaEqTrans} &
-\ttindexboldpos{==}{$IsaEq} &
-\verb$\<rightleftharpoons>$ \\
-\indexboldpos{\isasymrightharpoonup}{$IsaEqTrans1} &
-\ttindexboldpos{=>}{$IsaFun} &
-\verb$\<rightharpoonup>$ \\
-\indexboldpos{\isasymleftharpoondown}{$IsaEqTrans2} &
-\ttindexboldpos{<=}{$IsaFun2} &
-\verb$\<leftharpoondown>$ \\
-\indexboldpos{\isasymlambda}{$Isalam} &
-\texttt{\%}\indexbold{$Isalam@\texttt{\%}} &
-\verb$\<lambda>$ \\
-\indexboldpos{\isasymFun}{$IsaFun} &
-\ttindexboldpos{=>}{$IsaFun} &
-\verb$\<Rightarrow>$ \\
-\indexboldpos{\isasymand}{$HOL0and} &
-\texttt{\&}\indexbold{$HOL0and@{\texttt{\&}}} &
-\verb$\<and>$ \\
-\indexboldpos{\isasymor}{$HOL0or} &
-\texttt{|}\index{$HOL0or@\ttor|bold} &
-\verb$\<or>$ \\
-\indexboldpos{\isasymimp}{$HOL0imp} &
-\ttindexboldpos{-->}{$HOL0imp} &
-\verb$\<longrightarrow>$ \\
-\indexboldpos{\isasymnot}{$HOL0not} &
-\verb$~$\index{$HOL0not@\verb$~$|bold} &
-\verb$\<not>$ \\
-\indexboldpos{\isasymnoteq}{$HOL0noteq} &
-\verb$~=$\index{$HOL0noteq@\verb$~=$|bold} &
-\verb$\<noteq>$ \\
-\indexboldpos{\isasymforall}{$HOL0All} &
-\ttindexbold{ALL}, \texttt{!}\index{$HOL0All@\ttall|bold} &
-\verb$\<forall>$ \\
-\indexboldpos{\isasymexists}{$HOL0Ex} &
-\ttindexbold{EX}, \texttt{?}\index{$HOL0Ex@\texttt{?}|bold} &
-\verb$\<exists>$ \\
-\isasymuniqex\index{$HOL0ExU@\isasymuniqex|bold} &
-\ttEXU\index{EXX@\ttEXU|bold}, \ttuniquex\index{$HOL0ExU@\ttuniquex|bold} &
-\verb$\<exists>!$\\
-\indexboldpos{\isasymepsilon}{$HOL0ExSome} &
-\ttindexbold{SOME}, \isa{\at}\index{$HOL2list@\isa{\at}} &
-\verb$\<epsilon>$\\
-\indexboldpos{\isasymcirc}{$HOL1} &
-\ttindexbold{o} &
-\verb$\<circ>$\\
-\indexboldpos{\isasymbar~\isasymbar}{$HOL2arithfun}&
-\ttindexbold{abs}&
-\verb$\<bar> \<bar>$\\
-\indexboldpos{\isasymle}{$HOL2arithrel}&
-\isadxboldpos{<=}{$HOL2arithrel}&
-\verb$\<le>$\\
-\indexboldpos{\isasymtimes}{$Isatype}&
-\ttindexboldpos{*}{$HOL2arithfun} &
-\verb$\<times>$\\
-\indexboldpos{\isasymin}{$HOL3Set0a}&
-\ttindexboldpos{:}{$HOL3Set0b} &
-\verb$\<in>$\\
-\isasymnotin\index{$HOL3Set0c@\isasymnotin|bold} &
-\verb$~:$\index{$HOL3Set0d@\verb$~:$|bold} &
-\verb$\<notin>$\\
-\indexboldpos{\isasymsubseteq}{$HOL3Set0e}&
-\verb$<=$ & \verb$\<subseteq>$\\
-\indexboldpos{\isasymsubset}{$HOL3Set0f}&
-\verb$<$ & \verb$\<subset>$\\
-\indexboldpos{\isasymunion}{$HOL3Set1}&
-\ttindexbold{Un} &
-\verb$\<union>$\\
-\indexboldpos{\isasyminter}{$HOL3Set1}&
-\ttindexbold{Int} &
-\verb$\<inter>$\\
-\isasymUnion\index{$HOL3Set2@\isasymUnion|bold}&
-\ttindexbold{UN}, \ttindexbold{Union} &
-\verb$\<Union>$\\
-\isasymInter\index{$HOL3Set2@\isasymInter|bold}&
-\ttindexbold{INT}, \ttindexbold{Inter} &
-\verb$\<Inter>$\\
-\isactrlsup{\isacharasterisk}\index{$HOL4star@\isactrlsup{\isacharasterisk}|bold}&
-\verb$^*$\index{$HOL4star@\verb$^$\texttt{*}|bold} &
-\verb$\<^sup>*$\\
-\isasyminverse\index{$HOL4inv@\isasyminverse|bold}&
-\verb$^-1$\index{$HOL4inv@\verb$^-1$|bold} &
-\verb$\<inverse>$\\
-\hline
-\end{tabular}
-\end{center}
-\caption{Mathematical Symbols, Their \textsc{ascii}-Equivalents and Internal Names}
-\label{tab:ascii}
-\end{table}\indexbold{ASCII@\textsc{ascii} symbols}
-
-\input{document/appendix.tex}
-
-\begin{table}[htbp]
-\begin{center}
-\begin{tabular}{@{}|lllllllll|@{}}
-\hline
-\texttt{ALL} &
-\texttt{BIT} &
-\texttt{CHR} &
-\texttt{EX} &
-\texttt{GREATEST} &
-\texttt{INT} &
-\texttt{Int} &
-\texttt{LEAST} &
-\texttt{O} \\
-\texttt{OFCLASS} &
-\texttt{PI} &
-\texttt{PROP} &
-\texttt{SIGMA} &
-\texttt{SOME} &
-\texttt{THE} &
-\texttt{TYPE} &
-\texttt{UN} &
-\texttt{Un} \\
-\texttt{WRT} &
-\texttt{case} &
-\texttt{choose} &
-\texttt{div} &
-\texttt{dvd} &
-\texttt{else} &
-\texttt{funcset} &
-\texttt{if} &
-\texttt{in} \\
-\texttt{let} &
-\texttt{mem} &
-\texttt{mod} &
-\texttt{o} &
-\texttt{of} &
-\texttt{op} &
-\texttt{then} &&\\
-\hline
-\end{tabular}
-\end{center}
-\caption{Reserved Words in HOL Terms}
-\label{tab:ReservedWords}
-\end{table}
-
-
-%\begin{table}[htbp]
-%\begin{center}
-%\begin{tabular}{|lllll|}
-%\hline
-%\texttt{and} &
-%\texttt{binder} &
-%\texttt{concl} &
-%\texttt{congs} \\
-%\texttt{distinct} &
-%\texttt{files} &
-%\texttt{in} &
-%\texttt{induction} &
-%\texttt{infixl} \\
-%\texttt{infixr} &
-%\texttt{inject} &
-%\texttt{intrs} &
-%\texttt{is} &
-%\texttt{monos} \\
-%\texttt{output} &
-%\texttt{where} &
-% &
-% &
-% \\
-%\hline
-%\end{tabular}
-%\end{center}
-%\caption{Minor Keywords in HOL Theories}
-%\label{tab:keywords}
-%\end{table}
--- a/doc-src/TutorialI/basics.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,350 +0,0 @@
-\chapter{The Basics}
-
-\section{Introduction}
-
-This book is a tutorial on how to use the theorem prover Isabelle/HOL as a
-specification and verification system. Isabelle is a generic system for
-implementing logical formalisms, and Isabelle/HOL is the specialization
-of Isabelle for HOL, which abbreviates Higher-Order Logic. We introduce
-HOL step by step following the equation
-\[ \mbox{HOL} = \mbox{Functional Programming} + \mbox{Logic}. \]
-We do not assume that you are familiar with mathematical logic.
-However, we do assume that
-you are used to logical and set theoretic notation, as covered
-in a good discrete mathematics course~\cite{Rosen-DMA}, and
-that you are familiar with the basic concepts of functional
-programming~\cite{Bird-Haskell,Hudak-Haskell,paulson-ml2,Thompson-Haskell}.
-Although this tutorial initially concentrates on functional programming, do
-not be misled: HOL can express most mathematical concepts, and functional
-programming is just one particularly simple and ubiquitous instance.
-
-Isabelle~\cite{paulson-isa-book} is implemented in ML~\cite{SML}. This has
-influenced some of Isabelle/HOL's concrete syntax but is otherwise irrelevant
-for us: this tutorial is based on
-Isabelle/Isar~\cite{isabelle-isar-ref}, an extension of Isabelle which hides
-the implementation language almost completely. Thus the full name of the
-system should be Isabelle/Isar/HOL, but that is a bit of a mouthful.
-
-There are other implementations of HOL, in particular the one by Mike Gordon
-\index{Gordon, Mike}%
-\emph{et al.}, which is usually referred to as ``the HOL system''
-\cite{mgordon-hol}. For us, HOL refers to the logical system, and sometimes
-its incarnation Isabelle/HOL\@.
-
-A tutorial is by definition incomplete. Currently the tutorial only
-introduces the rudiments of Isar's proof language. To fully exploit the power
-of Isar, in particular the ability to write readable and structured proofs,
-you should start with Nipkow's overview~\cite{Nipkow-TYPES02} and consult
-the Isabelle/Isar Reference Manual~\cite{isabelle-isar-ref} and Wenzel's
-PhD thesis~\cite{Wenzel-PhD} (which discusses many proof patterns)
-for further details. If you want to use Isabelle's ML level
-directly (for example for writing your own proof procedures) see the Isabelle
-Reference Manual~\cite{isabelle-ref}; for details relating to HOL see the
-Isabelle/HOL manual~\cite{isabelle-HOL}. All manuals have a comprehensive
-index.
-
-\section{Theories}
-\label{sec:Basic:Theories}
-
-\index{theories|(}%
-Working with Isabelle means creating theories. Roughly speaking, a
-\textbf{theory} is a named collection of types, functions, and theorems,
-much like a module in a programming language or a specification in a
-specification language. In fact, theories in HOL can be either. The general
-format of a theory \texttt{T} is
-\begin{ttbox}
-theory T
-imports B\(@1\) \(\ldots\) B\(@n\)
-begin
-{\rmfamily\textit{declarations, definitions, and proofs}}
-end
-\end{ttbox}\cmmdx{theory}\cmmdx{imports}
-where \texttt{B}$@1$ \dots\ \texttt{B}$@n$ are the names of existing
-theories that \texttt{T} is based on and \textit{declarations,
- definitions, and proofs} represents the newly introduced concepts
-(types, functions etc.) and proofs about them. The \texttt{B}$@i$ are the
-direct \textbf{parent theories}\indexbold{parent theories} of~\texttt{T}\@.
-Everything defined in the parent theories (and their parents, recursively) is
-automatically visible. To avoid name clashes, identifiers can be
-\textbf{qualified}\indexbold{identifiers!qualified}
-by theory names as in \texttt{T.f} and~\texttt{B.f}.
-Each theory \texttt{T} must
-reside in a \textbf{theory file}\index{theory files} named \texttt{T.thy}.
-
-This tutorial is concerned with introducing you to the different linguistic
-constructs that can fill the \textit{declarations, definitions, and
- proofs} above. A complete grammar of the basic
-constructs is found in the Isabelle/Isar Reference
-Manual~\cite{isabelle-isar-ref}.
-
-\begin{warn}
- HOL contains a theory \thydx{Main}, the union of all the basic
- predefined theories like arithmetic, lists, sets, etc.
- Unless you know what you are doing, always include \isa{Main}
- as a direct or indirect parent of all your theories.
-\end{warn}
-HOL's theory collection is available online at
-\begin{center}\small
- \url{http://isabelle.in.tum.de/library/HOL/}
-\end{center}
-and is recommended browsing. In subdirectory \texttt{Library} you find
-a growing library of useful theories that are not part of \isa{Main}
-but can be included among the parents of a theory and will then be
-loaded automatically.
-
-For the more adventurous, there is the \emph{Archive of Formal Proofs},
-a journal-like collection of more advanced Isabelle theories:
-\begin{center}\small
- \url{http://afp.sourceforge.net/}
-\end{center}
-We hope that you will contribute to it yourself one day.%
-\index{theories|)}
-
-
-\section{Types, Terms and Formulae}
-\label{sec:TypesTermsForms}
-
-Embedded in a theory are the types, terms and formulae of HOL\@. HOL is a typed
-logic whose type system resembles that of functional programming languages
-like ML or Haskell. Thus there are
-\index{types|(}
-\begin{description}
-\item[base types,]
-in particular \tydx{bool}, the type of truth values,
-and \tydx{nat}, the type of natural numbers.
-\item[type constructors,]\index{type constructors}
- in particular \tydx{list}, the type of
-lists, and \tydx{set}, the type of sets. Type constructors are written
-postfix, e.g.\ \isa{(nat)list} is the type of lists whose elements are
-natural numbers. Parentheses around single arguments can be dropped (as in
-\isa{nat list}), multiple arguments are separated by commas (as in
-\isa{(bool,nat)ty}).
-\item[function types,]\index{function types}
-denoted by \isasymFun\indexbold{$IsaFun@\isasymFun}.
- In HOL \isasymFun\ represents \emph{total} functions only. As is customary,
- \isa{$\tau@1$ \isasymFun~$\tau@2$ \isasymFun~$\tau@3$} means
- \isa{$\tau@1$ \isasymFun~($\tau@2$ \isasymFun~$\tau@3$)}. Isabelle also
- supports the notation \isa{[$\tau@1,\dots,\tau@n$] \isasymFun~$\tau$}
- which abbreviates \isa{$\tau@1$ \isasymFun~$\cdots$ \isasymFun~$\tau@n$
- \isasymFun~$\tau$}.
-\item[type variables,]\index{type variables}\index{variables!type}
- denoted by \ttindexboldpos{'a}{$Isatype}, \isa{'b} etc., just like in ML\@. They give rise
- to polymorphic types like \isa{'a \isasymFun~'a}, the type of the identity
- function.
-\end{description}
-\begin{warn}
- Types are extremely important because they prevent us from writing
- nonsense. Isabelle insists that all terms and formulae must be
- well-typed and will print an error message if a type mismatch is
- encountered. To reduce the amount of explicit type information that
- needs to be provided by the user, Isabelle infers the type of all
- variables automatically (this is called \bfindex{type inference})
- and keeps quiet about it. Occasionally this may lead to
- misunderstandings between you and the system. If anything strange
- happens, we recommend that you ask Isabelle to display all type
- information via the Proof General menu item \pgmenu{Isabelle} $>$
- \pgmenu{Settings} $>$ \pgmenu{Show Types} (see \S\ref{sec:interface}
- for details).
-\end{warn}%
-\index{types|)}
-
-
-\index{terms|(}
-\textbf{Terms} are formed as in functional programming by
-applying functions to arguments. If \isa{f} is a function of type
-\isa{$\tau@1$ \isasymFun~$\tau@2$} and \isa{t} is a term of type
-$\tau@1$ then \isa{f~t} is a term of type $\tau@2$. HOL also supports
-infix functions like \isa{+} and some basic constructs from functional
-programming, such as conditional expressions:
-\begin{description}
-\item[\isa{if $b$ then $t@1$ else $t@2$}]\index{*if expressions}
-Here $b$ is of type \isa{bool} and $t@1$ and $t@2$ are of the same type.
-\item[\isa{let $x$ = $t$ in $u$}]\index{*let expressions}
-is equivalent to $u$ where all free occurrences of $x$ have been replaced by
-$t$. For example,
-\isa{let x = 0 in x+x} is equivalent to \isa{0+0}. Multiple bindings are separated
-by semicolons: \isa{let $x@1$ = $t@1$;\dots; $x@n$ = $t@n$ in $u$}.
-\item[\isa{case $e$ of $c@1$ \isasymFun~$e@1$ |~\dots~| $c@n$ \isasymFun~$e@n$}]
-\index{*case expressions}
-evaluates to $e@i$ if $e$ is of the form $c@i$.
-\end{description}
-
-Terms may also contain
-\isasymlambda-abstractions.\index{lambda@$\lambda$ expressions}
-For example,
-\isa{\isasymlambda{}x.~x+1} is the function that takes an argument \isa{x} and
-returns \isa{x+1}. Instead of
-\isa{\isasymlambda{}x.\isasymlambda{}y.\isasymlambda{}z.~$t$} we can write
-\isa{\isasymlambda{}x~y~z.~$t$}.%
-\index{terms|)}
-
-\index{formulae|(}%
-\textbf{Formulae} are terms of type \tydx{bool}.
-There are the basic constants \cdx{True} and \cdx{False} and
-the usual logical connectives (in decreasing order of priority):
-\indexboldpos{\protect\isasymnot}{$HOL0not}, \indexboldpos{\protect\isasymand}{$HOL0and},
-\indexboldpos{\protect\isasymor}{$HOL0or}, and \indexboldpos{\protect\isasymimp}{$HOL0imp},
-all of which (except the unary \isasymnot) associate to the right. In
-particular \isa{A \isasymimp~B \isasymimp~C} means \isa{A \isasymimp~(B
- \isasymimp~C)} and is thus logically equivalent to \isa{A \isasymand~B
- \isasymimp~C} (which is \isa{(A \isasymand~B) \isasymimp~C}).
-
-Equality\index{equality} is available in the form of the infix function
-\isa{=} of type \isa{'a \isasymFun~'a
- \isasymFun~bool}. Thus \isa{$t@1$ = $t@2$} is a formula provided $t@1$
-and $t@2$ are terms of the same type. If $t@1$ and $t@2$ are of type
-\isa{bool} then \isa{=} acts as \rmindex{if-and-only-if}.
-The formula
-\isa{$t@1$~\isasymnoteq~$t@2$} is merely an abbreviation for
-\isa{\isasymnot($t@1$ = $t@2$)}.
-
-Quantifiers\index{quantifiers} are written as
-\isa{\isasymforall{}x.~$P$} and \isa{\isasymexists{}x.~$P$}.
-There is even
-\isa{\isasymuniqex{}x.~$P$}, which
-means that there exists exactly one \isa{x} that satisfies \isa{$P$}.
-Nested quantifications can be abbreviated:
-\isa{\isasymforall{}x~y~z.~$P$} means
-\isa{\isasymforall{}x.\isasymforall{}y.\isasymforall{}z.~$P$}.%
-\index{formulae|)}
-
-Despite type inference, it is sometimes necessary to attach explicit
-\bfindex{type constraints} to a term. The syntax is
-\isa{$t$::$\tau$} as in \isa{x < (y::nat)}. Note that
-\ttindexboldpos{::}{$Isatype} binds weakly and should therefore be enclosed
-in parentheses. For instance,
-\isa{x < y::nat} is ill-typed because it is interpreted as
-\isa{(x < y)::nat}. Type constraints may be needed to disambiguate
-expressions
-involving overloaded functions such as~\isa{+},
-\isa{*} and~\isa{<}. Section~\ref{sec:overloading}
-discusses overloading, while Table~\ref{tab:overloading} presents the most
-important overloaded function symbols.
-
-In general, HOL's concrete \rmindex{syntax} tries to follow the conventions of
-functional programming and mathematics. Here are the main rules that you
-should be familiar with to avoid certain syntactic traps:
-\begin{itemize}
-\item
-Remember that \isa{f t u} means \isa{(f t) u} and not \isa{f(t u)}!
-\item
-Isabelle allows infix functions like \isa{+}. The prefix form of function
-application binds more strongly than anything else and hence \isa{f~x + y}
-means \isa{(f~x)~+~y} and not \isa{f(x+y)}.
-\item Remember that in HOL if-and-only-if is expressed using equality. But
- equality has a high priority, as befitting a relation, while if-and-only-if
- typically has the lowest priority. Thus, \isa{\isasymnot~\isasymnot~P =
- P} means \isa{\isasymnot\isasymnot(P = P)} and not
- \isa{(\isasymnot\isasymnot P) = P}. When using \isa{=} to mean
- logical equivalence, enclose both operands in parentheses, as in \isa{(A
- \isasymand~B) = (B \isasymand~A)}.
-\item
-Constructs with an opening but without a closing delimiter bind very weakly
-and should therefore be enclosed in parentheses if they appear in subterms, as
-in \isa{(\isasymlambda{}x.~x) = f}. This includes
-\isa{if},\index{*if expressions}
-\isa{let},\index{*let expressions}
-\isa{case},\index{*case expressions}
-\isa{\isasymlambda}, and quantifiers.
-\item
-Never write \isa{\isasymlambda{}x.x} or \isa{\isasymforall{}x.x=x}
-because \isa{x.x} is always taken as a single qualified identifier. Write
-\isa{\isasymlambda{}x.~x} and \isa{\isasymforall{}x.~x=x} instead.
-\item Identifiers\indexbold{identifiers} may contain the characters \isa{_}
-and~\isa{'}, except at the beginning.
-\end{itemize}
-
-For the sake of readability, we use the usual mathematical symbols throughout
-the tutorial. Their \textsc{ascii}-equivalents are shown in table~\ref{tab:ascii} in
-the appendix.
-
-\begin{warn}
-A particular problem for novices can be the priority of operators. If
-you are unsure, use additional parentheses. In those cases where
-Isabelle echoes your input, you can see which parentheses are dropped
---- they were superfluous. If you are unsure how to interpret
-Isabelle's output because you don't know where the (dropped)
-parentheses go, set the Proof General flag \pgmenu{Isabelle} $>$
-\pgmenu{Settings} $>$ \pgmenu{Show Brackets} (see \S\ref{sec:interface}).
-\end{warn}
-
-
-\section{Variables}
-\label{sec:variables}
-\index{variables|(}
-
-Isabelle distinguishes free and bound variables, as is customary. Bound
-variables are automatically renamed to avoid clashes with free variables. In
-addition, Isabelle has a third kind of variable, called a \textbf{schematic
- variable}\index{variables!schematic} or \textbf{unknown}\index{unknowns},
-which must have a~\isa{?} as its first character.
-Logically, an unknown is a free variable. But it may be
-instantiated by another term during the proof process. For example, the
-mathematical theorem $x = x$ is represented in Isabelle as \isa{?x = ?x},
-which means that Isabelle can instantiate it arbitrarily. This is in contrast
-to ordinary variables, which remain fixed. The programming language Prolog
-calls unknowns {\em logical\/} variables.
-
-Most of the time you can and should ignore unknowns and work with ordinary
-variables. Just don't be surprised that after you have finished the proof of
-a theorem, Isabelle will turn your free variables into unknowns. It
-indicates that Isabelle will automatically instantiate those unknowns
-suitably when the theorem is used in some other proof.
-Note that for readability we often drop the \isa{?}s when displaying a theorem.
-\begin{warn}
- For historical reasons, Isabelle accepts \isa{?} as an ASCII representation
- of the \(\exists\) symbol. However, the \isa{?} character must then be followed
- by a space, as in \isa{?~x. f(x) = 0}. Otherwise, \isa{?x} is
- interpreted as a schematic variable. The preferred ASCII representation of
- the \(\exists\) symbol is \isa{EX}\@.
-\end{warn}%
-\index{variables|)}
-
-\section{Interaction and Interfaces}
-\label{sec:interface}
-
-The recommended interface for Isabelle/Isar is the (X)Emacs-based
-\bfindex{Proof General}~\cite{proofgeneral,Aspinall:TACAS:2000}.
-Interaction with Isabelle at the shell level, although possible,
-should be avoided. Most of the tutorial is independent of the
-interface and is phrased in a neutral language. For example, the
-phrase ``to abandon a proof'' corresponds to the obvious
-action of clicking on the \pgmenu{Undo} symbol in Proof General.
-Proof General specific information is often displayed in paragraphs
-identified by a miniature Proof General icon. Here are two examples:
-\begin{pgnote}
-Proof General supports a special font with mathematical symbols known
-as ``x-symbols''. All symbols have \textsc{ascii}-equivalents: for
-example, you can enter either \verb!&! or \verb!\<and>! to obtain
-$\land$. For a list of the most frequent symbols see table~\ref{tab:ascii}
-in the appendix.
-
-Note that by default x-symbols are not enabled. You have to switch
-them on via the menu item \pgmenu{Proof-General} $>$ \pgmenu{Options} $>$
-\pgmenu{X-Symbols} (and save the option via the top-level
-\pgmenu{Options} menu).
-\end{pgnote}
-
-\begin{pgnote}
-Proof General offers the \pgmenu{Isabelle} menu for displaying
-information and setting flags. A particularly useful flag is
-\pgmenu{Isabelle} $>$ \pgmenu{Settings} $>$ \pgdx{Show Types} which
-causes Isabelle to output the type information that is usually
-suppressed. This is indispensible in case of errors of all kinds
-because often the types reveal the source of the problem. Once you
-have diagnosed the problem you may no longer want to see the types
-because they clutter all output. Simply reset the flag.
-\end{pgnote}
-
-\section{Getting Started}
-
-Assuming you have installed Isabelle and Proof General, you start it by typing
-\texttt{Isabelle} in a shell window. This launches a Proof General window.
-By default, you are in HOL\footnote{This is controlled by the
-\texttt{ISABELLE_LOGIC} setting, see \emph{The Isabelle System Manual}
-for more details.}.
-
-\begin{pgnote}
-You can choose a different logic via the \pgmenu{Isabelle} $>$
-\pgmenu{Logics} menu.
-\end{pgnote}
--- a/doc-src/TutorialI/cl2emono-modified.sty Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1371 +0,0 @@
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%% This is cl2emono.sty version 2.2
-%% (intermediate fix also for springer.sty for the use of
-%% LaTeX2e and NFSS2)
-%%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is ucgreek
-% the definition of versal greek characters
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\mathchardef\Gamma="0100
-\mathchardef\Delta="0101
-\mathchardef\Theta="0102
-\mathchardef\Lambda="0103
-\mathchardef\Xi="0104
-\mathchardef\Pi="0105
-\mathchardef\Sigma="0106
-\mathchardef\Upsilon="0107
-\mathchardef\Phi="0108
-\mathchardef\Psi="0109
-\mathchardef\Omega="010A
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is referee.tex
-%
-% It defines the style option "referee"
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\newif\if@referee \@refereefalse
-\def\ds@referee{\@refereetrue
-\typeout{A referee's copy will be produced.}%
-\def\baselinestretch{2}\small
-\normalsize\rm
-\newbox\refereebox
-\setbox\refereebox=\vbox to0pt{\vskip0.5cm%
- \hbox to\textwidth{\normalsize\tt\hrulefill\lower0.5ex
- \hbox{\kern5pt referee's copy\kern5pt}\hrulefill}\vss}%
-\def\@oddfoot{\copy\refereebox}\let\@evenfoot=\@oddfoot}
-% This is footinfo.tex
-% it provides an informatory line on every page
-%
-\def\SpringerMacroPackageNameA{\springerstylefile}
-% \thetime, \thedate and \timstamp are macros to include
-% time, date (or both) of the TeX run in the document
-\def\maketimestamp{\count255=\time
-\divide\count255 by 60\relax
-\edef\thetime{\the\count255:}%
-\multiply\count255 by-60\relax
-\advance\count255 by\time
-\edef\thetime{\thetime\ifnum\count255<10 0\fi\the\count255}
-\edef\thedate{\number\day-\ifcase\month\or Jan\or Feb\or Mar\or
- Apr\or May\or Jun\or Jul\or Aug\or Sep\or Oct\or
- Nov\or Dec\fi-\number\year}
-\def\timstamp{\hbox to\hsize{\tt\hfil\thedate\hfil\thetime\hfil}}}
-\maketimestamp
-%
-% \footinfo generates a info footline on every page containing
-% pagenumber, jobname, macroname, and timestamp
-\def\ds@footinfo{\maketimestamp
- \def\@oddfoot{\footnotesize\tt Page: \thepage\hfil job: \jobname\hfil
- macro: \SpringerMacroPackageNameA\hfil
- date/time: \thedate/\thetime}%
- \let\@evenfoot=\@oddfoot}
-\def\footinfo{\maketimestamp
- \ds@footinfo
- \typeout{You ordered a foot-info line. }}
-\def\nofootinfo{%
- \def\@oddfoot{}\def\@evenfoot{}%
- \typeout{Foot-info has been disabled. }}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is vector.tex
-%
-% It redefines the plain TeX \vec command
-% and adds a \tens command for tensors
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-% ##### (changed by RB)
-\def\vec@style{\relax} % hook to change style of vector
- % default yields boldface italic
-% \def\vec@style{\bf} % hook to change style of vector
-% % default yields mathstyle i.e. boldface upright
-
-\def\vec#1{\relax\ifmmode\mathchoice
-{\mbox{\boldmath$\vec@style\displaystyle#1$}}
-{\mbox{\boldmath$\vec@style\textstyle#1$}}
-{\mbox{\boldmath$\vec@style\scriptstyle#1$}}
-{\mbox{\boldmath$\vec@style\scriptscriptstyle#1$}}\else
-\hbox{\boldmath$\vec@style\textstyle#1$}\fi}
-
-\def\tens#1{\relax\ifmmode\mathchoice{\mbox{$\sf\displaystyle#1$}}
-{\mbox{$\sf\textstyle#1$}}
-{\mbox{$\sf\scriptstyle#1$}}
-{\mbox{$\sf\scriptscriptstyle#1$}}\else
-\hbox{$\sf\textstyle#1$}\fi}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is vecstyle.tex
-%
-% It enables documentstyle options vecmath and vecphys
-% to change the vectors to upright bold face or
-% to math italic bold respectively
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\ds@vecmath{\def\vec@style{\bf}\typeout{Vectors are now represented
-in mathematics style as boldface upright characters.}}
-\def\ds@vecphys{\let\vec@style\relax\typeout{Vectors are now represented
-in physics style as boldface italics.}}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is defaults.tex
-%
-% It sets the switches for twoside printing, numbering
-% of equations and kind of citation.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\@twosidetrue % twoside is default
-\newif\if@bibay \@bibayfalse % citation with numbers
- % is default
-\newif\if@numart \@numartfalse % numbering with chapternumbers
- % is default
-
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is misc.xxx
-%
-% It defines various commands not available in "plain LaTeX"
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\newcommand{\ts}{\thinspace{}}
-\newcommand{\sq}{\hbox{\rlap{$\sqcap$}$\sqcup$}}
-\newcommand{\qed}{\ifmmode\sq\else{\unskip\nobreak\hfil
- \penalty50\hskip1em\null\nobreak\hfil\sq
- \parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi{}}
-\def\D{{\rm d}}
-\def\E{{\rm e}}
-\let\eul=\E
-\def\I{{\rm i}}
-\let\imag=\I
-\def\strich{\vskip0.5cm\hrule\vskip3ptplus12pt\null}
-
-% Frame for paste-in figures or tables
-%\def\mpicplace#1#2{%#1 = width #2 = height
-%\vbox{\@tempdima=#2\advance\@tempdima by-2\fboxrule
-%\hrule\@height \fboxrule\@width #1
-%\hbox to #1{\vrule\@width \fboxrule\@height\@tempdima\hfil
-%\vrule\@width \fboxrule\@height\@tempdima}\hrule\@height
-%\fboxrule\@width #1}}
-
-% #####
-% Frame for paste-in figures or tables
-\def\mpicplace#1#2{% #1 =width #2 =height
-\vbox{\hbox to #1{\vrule width \fboxrule height #2\hfill}}}
-
-\def\picplace#1{\mpicplace{\hsize}{#1}}
-% Ragged bottom for the actual page
-\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
-\global\let\@textbottom\relax}}
-\flushbottom
-
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is layout.m01
-%
-% It defines various sizes and settings for books
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\topmargin=0cm
-\textwidth=11.7cm
-\textheight=18.9cm
-\oddsidemargin=0.7cm
-\evensidemargin=0.7cm
-\headsep=12pt
-
-\baselineskip=12pt
-\parindent=15pt
-\parskip=0pt plus 1pt
-\hfuzz=2pt
-\frenchspacing
-
-\tolerance=500
-
-\abovedisplayskip=3mm plus6pt minus 4pt
-\belowdisplayskip=3mm plus6pt minus 4pt
-\abovedisplayshortskip=0mm plus6pt minus 2pt
-\belowdisplayshortskip=2mm plus4pt minus 4pt
-
-\predisplaypenalty=0
-\clubpenalty=10000
-\widowpenalty=10000
-
-\newdimen\betweenumberspace % dimension for space between
-\betweenumberspace=5pt % number and text of titles.
-\newdimen\headlineindent % dimension for space between
-\headlineindent=2.5cc % number and text of headings.
-
-\intextsep 20pt plus 2pt minus 2pt
-
-\setcounter{topnumber}{4}
-\def\topfraction{.9}
-\setcounter{bottomnumber}{2}
-\def\bottomfraction{.5}
-\setcounter{totalnumber}{6}
-\def\textfraction{.2}
-\def\floatpagefraction{.5}
-
-% Figures and tables are processed in small print
-\def \@floatboxreset {%
- \reset@font
- \small
- \@setnobreak
- \@setminipage
-}
-\def\figure{\@float{figure}}
-\@namedef{figure*}{\@dblfloat{figure}}
-\def\table{\@float{table}}
-\@namedef{table*}{\@dblfloat{table}}
-\def\fps@figure{htbp}
-\def\fps@table{htbp}
-
-\labelsep=5\p@ % measures for lists
-\leftmargini=17.777\p@
-\labelwidth\leftmargini\advance\labelwidth-\labelsep
-\leftmarginii=\leftmargini
-\leftmarginiii=\leftmargini
-\parsep=\parskip
-
-\def\@listI{\leftmargin\leftmargini
- \parsep=\parskip
- \topsep=\medskipamount
- \itemsep=\parskip \advance\itemsep by -\parsep}
-\let\@listi\@listI
-\@listi
-
-\def\@listii{\leftmargin\leftmarginii
- \labelwidth\leftmarginii\advance\labelwidth by -\labelsep
- \parsep=\parskip
- \topsep=\z@
- \itemsep=\parskip \advance\itemsep by -\parsep}
-\def\@listiii{\leftmargin\leftmarginiii
- \labelwidth\leftmarginiii\advance\labelwidth by -\labelsep
- \parsep=\parskip
- \topsep=\z@
- \itemsep=\parskip \advance\itemsep by -\parsep}
-%
-\def\@normalsize{\@setsize\normalsize{12pt}\xpt\@xpt
-\abovedisplayskip=3mm plus6pt minus 4pt
-\belowdisplayskip=3mm plus6pt minus 4pt
-\abovedisplayshortskip=0mm plus6pt minus 2pt
-\belowdisplayshortskip=2mm plus4pt minus 4pt
-\let\@listi\@listI} % Setting of \@listi added 9 Jun 87
-%
-\def\small{\@setsize\small{10pt}\ixpt\@ixpt
-\abovedisplayskip 8.5\p@ plus3\p@ minus4\p@
-\belowdisplayskip \abovedisplayskip
-\abovedisplayshortskip \z@ plus2\p@
-\belowdisplayshortskip 4\p@ plus2\p@ minus2\p@
-\def\@listi{\leftmargin\leftmargini
-\topsep 4pt plus 2pt minus 2pt\parsep\parskip
-\itemsep\parskip}}
-%
-\def\itemize{\ifnum \@itemdepth >3 \@toodeep\else \advance\@itemdepth \@ne
-\ifnum \@itemdepth=1 \leftmargini=10\p@
-\labelwidth\leftmargini\advance\labelwidth-\labelsep
-\leftmarginii=\leftmargini \leftmarginiii=\leftmargini \fi
-\edef\@itemitem{labelitem\romannumeral\the\@itemdepth}%
-\list{\csname\@itemitem\endcsname}{\def\makelabel##1{\rlap{##1}\hss}}\fi}
-%
-\newdimen\verbatimindent \verbatimindent\parindent
-\def\verbatim{\advance\@totalleftmargin by\verbatimindent
-\@verbatim \frenchspacing\@vobeyspaces \@xverbatim}
-%
-\let\footnotesize=\small
-%
-\def\petit{\par\addvspace{6pt}\small}
-\def\endpetit{\par\addvspace{6pt}}
-%
-\raggedbottom
-\normalsize % Choose the normalsize font.
-% This is texte.tex
-% it defines various texts and their translations
-% called up with documentstyle options
-\def\abstractname{Summary.}
-\def\ackname{Acknowledgement.}
-\def\andname{and}
-\def\lastandname{, and}
-\def\appendixname{Appendix}
-\def\chaptername{Chapter}
-\def\claimname{Claim}
-\def\conjecturename{Conjecture}
-\def\contentsname{Table of Contents}
-\def\corollaryname{Corollary}
-\def\definitionname{Definition}
-\def\examplename{Example}
-\def\exercisename{Exercise}
-\def\figurename{Fig.}
-\def\keywordname{{\bf Key words:}}
-\def\indexname{Index}
-\def\lemmaname{Lemma}
-\def\contriblistname{List of Contributors}
-\def\listfigurename{List of Figures}
-\def\listtablename{List of Tables}
-\def\mailname{{\it Correspondence to\/}:}
-\def\noteaddname{Note added in proof}
-\def\notename{Note}
-\def\partname{Part}
-\def\problemname{Problem}
-\def\proofname{Proof}
-\def\propertyname{Property}
-\def\propositionname{Proposition}
-\def\questionname{Question}
-\def\remarkname{Remark}
-\def\seename{see}
-\def\solutionname{Solution}
-\def\subclassname{{\it Subject Classifications\/}:}
-\def\tablename{Table}
-\def\theoremname{Theorem}
-% Names of theorem like environments are already defined
-% but must be translated if another language is chosen
-%
-% French section
-\def\ds@francais{\typeout{On parle francais.}%
- \def\abstractname{R\'esum\'e.}%
- \def\ackname{Remerciements.}%
- \def\andname{et}%
- \def\lastandname{ et}%
- \def\appendixname{Appendice}
- \def\chaptername{Chapitre}%
- \def\claimname{Pr\'etention}%
- \def\conjecturename{Hypoth\`ese}%
- \def\contentsname{Table des mati\`eres}%
- \def\corollaryname{Corollaire}%
- \def\definitionname{D\'efinition}%
- \def\examplename{Exemple}%
- \def\exercisename{Exercice}%
- \def\figurename{Fig.}%
- \def\keywordname{{\bf Mots-cl\'e:}}
- \def\indexname{Index}
- \def\lemmaname{Lemme}%
- \def\contriblistname{Liste des contributeurs}
- \def\listfigurename{Liste des figures}%
- \def\listtablename{Liste des tables}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
- \def\notename{Remarque}%
- \def\partname{Partie}%
- \def\problemname{Probl\`eme}%
- \def\proofname{\'Epreuve}%
- \def\propertyname{Caract\'eristique}%
-%\def\propositionname{Proposition}%
- \def\questionname{Question}%
- \def\remarkname{Remarque}%
- \def\seename{voir}
- \def\solutionname{Solution}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tableau}%
- \def\theoremname{Th\'eor\`eme}%
-}
-%
-% German section
-\def\ds@deutsch{\typeout{Man spricht deutsch.}%
- \def\abstractname{Zusammenfassung.}%
- \def\ackname{Danksagung.}%
- \def\andname{und}%
- \def\lastandname{ und}%
- \def\appendixname{Anhang}%
- \def\chaptername{Kapitel}%
- \def\claimname{Behauptung}%
- \def\conjecturename{Hypothese}%
- \def\contentsname{Inhaltsverzeichnis}%
- \def\corollaryname{Korollar}%
-%\def\definitionname{Definition}%
- \def\examplename{Beispiel}%
- \def\exercisename{\"Ubung}%
- \def\figurename{Abb.}%
- \def\keywordname{{\bf Schl\"usselw\"orter:}}
- \def\indexname{Index}
-%\def\lemmaname{Lemma}%
- \def\contriblistname{Mitarbeiter}
- \def\listfigurename{Abbildungsverzeichnis}%
- \def\listtablename{Tabellenverzeichnis}%
- \def\mailname{{\it Correspondence to\/}:}
- \def\noteaddname{Nachtrag}%
- \def\notename{Anmerkung}%
- \def\partname{Teil}%
-%\def\problemname{Problem}%
- \def\proofname{Beweis}%
- \def\propertyname{Eigenschaft}%
-%\def\propositionname{Proposition}%
- \def\questionname{Frage}%
- \def\remarkname{Anmerkung}%
- \def\seename{siehe}
- \def\solutionname{L\"osung}%
- \def\subclassname{{\it Subject Classifications\/}:}
- \def\tablename{Tabelle}%
-%\def\theoremname{Theorem}%
-}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is titneu.xxx
-%
-% It redefines titles. Usage is like Lamport described.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\setcounter{secnumdepth}{2} % depth of the highest-level
- % sectioning command
-\newcounter{chapter} % to use chapter together with
-\@addtoreset{section}{chapter} % article.sty
-\@addtoreset{footnote}{chapter}
-
-\def\thechapter{\arabic{chapter}} % how titles will be typeset
-\def\thesection{\thechapter.\arabic{section}}
-\def\thesubsection{\thesection.\arabic{subsection}}
-\def\thesubsubsection{\thesubsection.\arabic{subsubsection}}
-\def\theparagraph{\thesubsubsection.\arabic{paragraph}}
-\def\thesubparagraph{\theparagraph.\arabic{subparagraph}}
-\def\chaptermark#1{}
-\def\sec@hangfrom#1{\setbox\@tempboxa\hbox{#1}%
- \hangindent \z@\noindent\box\@tempboxa}
-
-% definition of chapter
-
-\def\@chapapp{\chaptername}
-
-\def\@makechapterhead#1{{\parindent0pt\raggedright
- \hyphenpenalty \@M
- \Large\bf\boldmath
- \sec@hangfrom{\thechapter\thechapterend\hskip\betweenumberspace}%!!!
- \ignorespaces#1\par
- \ifdim\pagetotal>118pt
- \vskip 24pt
- \else
- \@tempdima=118pt\advance\@tempdima by-\pagetotal
- \vskip\@tempdima
- \fi}}
-
-\def\@makeschapterhead#1{{\parindent0pt\raggedright
- \hyphenpenalty \@M
- \Large\bf\boldmath
- \ignorespaces#1\par
- \ifdim\pagetotal>118pt
- \vskip 24pt
- \else
- \@tempdima=118pt\advance\@tempdima by-\pagetotal
- \vskip\@tempdima
- \fi}}
-
-\newcommand{\clearemptydoublepage}{%
- \newpage{\pagestyle{empty}\cleardoublepage}}
-
-\def\chapter{\clearemptydoublepage\thispagestyle{empty}
- \global\@topnum\z@\@afterindentfalse
- \secdef\@chapter\@schapter}
-
-\def\@chapter[#1]#2{\ifnum\c@secnumdepth>\m@ne
- \refstepcounter{chapter}
- \typeout{\@chapapp\space\thechapter}
- \addcontentsline{toc}{chapter}{\protect
- \numberline{\thechapter\thechapterend}#1}\else %!!!
- \addcontentsline{toc}{chapter}{#1}
- \fi
- \chaptermark{#1}
- \addtocontents{lof}{\protect\addvspace{10pt}}
- \addtocontents{lot}{\protect\addvspace{10pt}}
- \if@twocolumn
- \@topnewpage[\@makechapterhead{#2}]
- \else \@makechapterhead{#2}
- \@afterheading
- \fi}
-
-\def\@schapter#1{\if@twocolumn\@topnewpage[\@makeschapterhead{#1}]
- \else \@makeschapterhead{#1}
- \@afterheading\fi}
-
-% Appendix
-\def\appendix{\par
- \setcounter{chapter}{0}%
- \setcounter{section}{0}%
- \def\@chapapp{\appendixname}%
- \def\thechapter{\Alph{chapter}}}
-
-% definition of sections
-% \hyphenpenalty and \raggedright added, so that there is no
-% hyphenation and the text is set ragged-right in sectioning
-
-\def\runinsep{.}
-\def\aftertext{\unskip\runinsep}
-
-\def\@sect#1#2#3#4#5#6[#7]#8{\ifnum #2>\c@secnumdepth
- \let\@svsec\@empty\else
- \refstepcounter{#1}\edef\@svsec{\csname the#1\endcsname
- \hskip\betweenumberspace
- \ignorespaces}\fi
- \@tempskipa #5\relax
- \ifdim \@tempskipa>\z@
- \begingroup #6\relax
- \sec@hangfrom{\hskip #3\relax\@svsec}{%
- \raggedright
- \hyphenpenalty \@M
- \interlinepenalty \@M #8\par}%
- \endgroup
- \csname #1mark\endcsname{#7}\addcontentsline
- {toc}{#1}{\ifnum #2>\c@secnumdepth \else
- \protect\numberline{\csname the#1\endcsname}\fi
- #7}\else
- \def\@svsechd{#6\hskip #3\relax
- \@svsec #8\aftertext\ignorespaces
- \csname #1mark\endcsname
- {#7}\addcontentsline
- {toc}{#1}{\ifnum #2>\c@secnumdepth \else
- \protect\numberline{\csname the#1\endcsname}\fi
- #7}}\fi
- \@xsect{#5}}
-
-% measures and setting of sections
-
-\def\section{\@startsection{section}{1}{\z@}%
- {-25pt plus-4pt minus-4pt}{12.5pt plus4pt
- minus4pt}{\large\bf\boldmath}}
-\def\subsection{\@startsection{subsection}{2}{\z@}%
- {-17pt plus-4pt minus-4pt}{10pt plus4pt
- minus4pt}{\normalsize\bf\boldmath}}
-\def\subsubsection{\@startsection{subsubsection}{3}{\z@}%
- {-5.388pt plus-4pt minus-4pt}{-5pt}{\normalsize\bf\boldmath}}
-\def\paragraph{\@startsection{paragraph}{4}{\z@}%
- {-5.388pt plus-4pt minus-4pt}{-5pt}{\normalsize\it}}
-\def\subparagraph{\@startsection{subparagraph}{5}{\z@}%
- {-5.388pt plus-4pt minus-4pt}{-5pt}{\normalsize\it}}
-
-% definition of \part
-\def\thepart{\Roman{part}}
-\def\part{\clearemptydoublepage % Starts new page.
- \thispagestyle{empty} % Page style of part page is empty
- \if@twocolumn % IF two-column style
- \onecolumn % THEN \onecolumn
- \@tempswatrue % @tempswa := true
- \else \@tempswafalse % ELSE @tempswa := false
- \fi % FI
-% \hbox{}\vfil % Add fil glue to center title
-%% \bgroup \centering % BEGIN centering %% Removed 19 Jan 88
- \secdef\@part\@spart}
-
-
-\def\@part[#1]#2{\ifnum \c@secnumdepth >-2\relax % IF secnumdepth > -2
- \refstepcounter{part} % THEN step part counter
- \addcontentsline{toc}{part}{\partname\ % add toc line
- \thepart. #1}\else % ELSE add unnumbered line
- \addcontentsline{toc}{part}{#1}\fi % FI
- \markboth{}{}
- {\raggedleft % added 8.1.92 FUH
- \ifnum \c@secnumdepth >-2\relax % IF secnumdepth > -2
- \Large\partname\ \thepart % THEN Print 'Part' and number
- \par % in \Large
- \vskip 103.3pt \fi % Add space before title.
- \bf\boldmath % FI
- #2\par}\@endpart} % Print Title in \Large bold.
-
-
-% \@endpart finishes the part page
-%
-\def\@endpart{\vfil\newpage % End page with 1fil glue.
- \if@twoside % IF twoside printing
- \hbox{} % THEN Produce totally blank page
- \thispagestyle{empty}
- \newpage
- \fi % FI
- \if@tempswa % IF @tempswa = true
- \twocolumn % THEN \twocolumn
- \fi} % FI
-
-\def\@spart#1{{\raggedleft % added 8 Jan 92 FUH
- \Large\bf\boldmath % Print title in \Large-boldface
- #1\par}\@endpart}
-
-\def\subtitle#1{\gdef\@subtitle{#1}}
-\def\@subtitle{}
-
-\def\maketitle{\par
- \begingroup
- \def\thefootnote{\fnsymbol{footnote}}%
- \def\@makefnmark{\hbox
- to\z@{$\m@th^{\@thefnmark}$\hss}}%
- \if@twocolumn
- \twocolumn[\@maketitle]%
- \else \newpage
- \global\@topnum\z@ % Prevents figures from going at top of page.
- \@maketitle \fi\thispagestyle{empty}\@thanks
- \par\penalty -\@M
- \endgroup
- \setcounter{footnote}{0}%
- \let\maketitle\relax
- \let\@maketitle\relax
- \gdef\@thanks{}\gdef\@author{}\gdef\@title{}\let\thanks\relax}
-
-\def\@maketitle{\newpage
- \null
- \vskip 2em % Vertical space above title.
-\begingroup
- \def\and{\unskip, }
- \parindent=\z@
- \pretolerance=10000
- \rightskip=0pt plus 3cm
- {\LARGE % each author set in \LARGE
- \lineskip .5em
- \@author
- \par}%
- \vskip 2cm % Vertical space after author.
- {\Huge \@title \par}% % Title set in \Huge size.
- \vskip 1cm % Vertical space after title.
- \if!\@subtitle!\else
- {\LARGE\ignorespaces\@subtitle \par}
- \vskip 1cm % Vertical space after subtitle.
- \fi
- \if!\@date!\else
- {\large \@date}% % Date set in \large size.
- \par
- \vskip 1.5em % Vertical space after date.
- \fi
- \vfill
- {\Large Springer-\kern-0.1em Verlag\par}
- \vskip 5pt
- \large
- Berlin\enspace Heidelberg\enspace New\kern0.1em York\\
- London\enspace Paris\enspace Tokyo\\
- Hong\thinspace Kong\enspace Barcelona\\
- Budapest\par
-\endgroup}
-
-\def\abstract{\if@twocolumn
-\section*{\abstractname}%
-\else \small
-\begin{center}%
-{\bf \abstractname\vspace{-.5em}\vspace{\z@}}%
-\end{center}%
-\quotation
-\fi}
-
-\def\endabstract{\if@twocolumn\else\endquotation\fi}
-
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is toc.xxx
-%
-% it modifies the appearence of the table of contents
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\def\tableofcontents{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\chapter*{\contentsname \@mkboth{{\contentsname}}{{\contentsname}}}
- \@starttoc{toc}\if@restonecol\twocolumn\fi}
-
-\setcounter{tocdepth}{2}
-
-\def\l@part#1#2{\addpenalty{\@secpenalty}%
- \addvspace{2em plus\p@}% % space above part line
- \begingroup
- \parindent \z@
- \rightskip \z@ plus 5em
- \hrule\vskip5pt
- \bf\boldmath % set line in boldface
- \leavevmode % TeX command to enter horizontal mode.
- #1\par
- \vskip5pt
- \hrule
- \vskip1pt
- \nobreak % Never break after part entry
- \endgroup}
-
-\def\@dotsep{2}
-
-\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
- \vskip 1.0em plus 1pt \@tempdima \tocchpnum \begingroup
- \parindent \z@ \rightskip \@pnumwidth
- \parfillskip -\@pnumwidth
- \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
- {\bf\boldmath#1}\nobreak
- \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
- \@dotsep mu$}\hfill
- \nobreak\hbox to\@pnumwidth{\hss #2}\par
- \penalty\@highpenalty \endgroup}
-
-\newdimen\tocchpnum
-\newdimen\tocsecnum
-\newdimen\tocsectotal
-\newdimen\tocsubsecnum
-\newdimen\tocsubsectotal
-\newdimen\tocsubsubsecnum
-\newdimen\tocsubsubsectotal
-\newdimen\tocparanum
-\newdimen\tocparatotal
-\newdimen\tocsubparanum
-\tocchpnum=20\p@ % chapter {\bf 88.} plus 5.3pt
-\tocsecnum=22.5\p@ % section 88.8. plus 4.722pt
-\tocsubsecnum=30.5\p@ % subsection 88.8.8 plus 4.944pt
-\tocsubsubsecnum=38\p@ % subsubsection 88.8.8.8 plus 4.666pt
-\tocparanum=45\p@ % paragraph 88.8.8.8.8 plus 3.888pt
-\tocsubparanum=53\p@ % subparagraph 88.8.8.8.8.8 plus 4.11pt
-\def\calctocindent{%
-\tocsectotal=\tocchpnum
-\advance\tocsectotal by\tocsecnum
-\tocsubsectotal=\tocsectotal
-\advance\tocsubsectotal by\tocsubsecnum
-\tocsubsubsectotal=\tocsubsectotal
-\advance\tocsubsubsectotal by\tocsubsubsecnum
-\tocparatotal=\tocsubsubsectotal
-\advance\tocparatotal by\tocparanum}
-\calctocindent
-
-\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
-\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
-\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
-\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
-\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
-
-\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\chapter*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
- \@starttoc{lof}\if@restonecol\twocolumn\fi}
-\def\l@figure{\@dottedtocline{1}{0em}{\tocsecnum}}
-
-\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
- \fi\chapter*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
- \@starttoc{lot}\if@restonecol\twocolumn\fi}
-\let\l@table\l@figure
-
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is runnhead.xxx
-%
-% It redefines the headings of a text. There are two
-% pagestyles possible: "\pagestyle{headings}" and
-% "\pagestyle{myheadings}". "\pagestyle{headings}" is
-% default.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-\@ifundefined{thechapterend}{\def\thechapterend{.}}{}
-\if@twoside
-\def\ps@headings{\let\@mkboth\markboth
- \def\@oddfoot{}\def\@evenfoot{}
- \def\@evenhead{\small\rm\rlap{\thepage}\hskip\headlineindent
- \leftmark\hfil}
- \def\@oddhead{\hfil\small\rm\rightmark\hskip\headlineindent
- \llap{\thepage}}
- \def\chaptermark##1{\markboth{{\ifnum\c@secnumdepth>\m@ne
- \thechapter\thechapterend\hskip\betweenumberspace\fi ##1}}{{\ifnum %!!!
- \c@secnumdepth>\m@ne\thechapter\thechapterend\hskip\betweenumberspace\fi ##1}}}%!!!
- \def\sectionmark##1{\markright{{\ifnum\c@secnumdepth>\z@
- \thesection\hskip\betweenumberspace\fi ##1}}}}
-\else \def\ps@headings{\let\@mkboth\markboth
- \def\@oddfoot{}\def\@evenfoot{}
- \def\@oddhead{\hfil\small\rm\rightmark\hskip\headlineindent
- \llap{\thepage}}
- \def\chaptermark##1{\markright{{\ifnum\c@secnumdepth>\m@ne
- \thechapter\thechapterend\hskip\betweenumberspace\fi ##1}}}} %!!!
-\fi
-\def\ps@myheadings{\let\@mkboth\@gobbletwo
- \def\@oddfoot{}\def\@evenfoot{}
- \def\@evenhead{\small\rm\rlap{\thepage}\hskip\headlineindent
- \leftmark\hfil}
- \def\@oddhead{\hfil\small\rm\rightmark\hskip\headlineindent
- \llap{\thepage}}
- \def\chaptermark##1{}
- \def\sectionmark##1{}%
- \def\subsectionmark##1{}}
-\ps@headings
-
-% Definition of the "\spnewtheorem" command.
-%
-% Usage:
-%
-% \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
-% or \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
-% or \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
-%
-% New is "cap_font" and "body_font". It stands for
-% fontdefinition of the caption and the text itself.
-%
-% "\spnewtheorem*" gives a theorem without number.
-%
-% A defined spnewthoerem environment is used as described
-% by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\let\if@envcntreset\iffalse % environment counter is reset each chapter
-\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
-\let\if@envcntsame\iffalse % NOT all environments like "Theorem",
- % each using its own counter
-\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
-\def\envankh{section} % show \thesection along with theorem number
-\DeclareOption{envcountchap}{\def\envankh{chapter}%
-\ExecuteOptions{envcountsect}}
-\let\if@envcntsect\iftrue % show \csname the\envankh\endcsname along
- % with environment number
-\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
-\ProcessOptions
-
-\def\@thmcountersep{.}
-\def\@thmcounterend{.}
-
-\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
-
-% definition of \spnewtheorem with number
-
-\def\@spnthm#1#2{%
- \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
-\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
-
-\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}\@addtoreset{#1}{#3}%
- \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
- \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\@definecounter{#1}%
- \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@spothm#1[#2]#3#4#5{%
- \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
- {\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{the#1}{\@nameuse{the#2}}%
- \expandafter\xdef\csname #1name\endcsname{#3}%
- \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
- \global\@namedef{end#1}{\@endtheorem}}}}
-
-\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\refstepcounter{#1}%
-\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
-
-\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
- \ignorespaces}
-
-\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
- the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
-
-\def\@spbegintheorem#1#2#3#4{\trivlist
- \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
-
-\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
- \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
-
-% definition of \spnewtheorem* without number
-
-\def\@sthm#1#2{\@Ynthm{#1}{#2}}
-
-\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
- {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
- \expandafter\xdef\csname #1name\endcsname{#2}%
- \global\@namedef{end#1}{\@endtheorem}}}
-
-\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
-\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
-
-\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
-
-\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
- {#4}{#2}{#3}\ignorespaces}
-
-\def\@Begintheorem#1#2#3{#3\trivlist
- \item[\hskip\labelsep{#2#1\@thmcounterend}]}
-
-\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
- \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
-
-% initialize theorem environment
-
-\if@envcntsect % show section counter
- \def\@thmcountersep{.}
- \spnewtheorem{theorem}{Theorem}[\envankh]{\bfseries}{\itshape}
-\else % theorem counter only
- \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
- \if@envcntreset
- \@addtoreset{theorem}{section}
- \else
- \@addtoreset{theorem}{chapter}
- \fi
-\fi
-
-%definition of divers theorem environments
-\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
-\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
-\if@envcntsame % all environments like "Theorem" - using its counter
- \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
-\else % all environments with their own counter
- \if@envcntsect % show section counter
- \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[\envankh]{#3}{#4}}
- \else % environment counter only
- \if@envcntreset % environment counter is reset each section
- \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
- \@addtoreset{#1}{section}}
- \else
- \let\spn@wtheorem=\@spynthm
- \fi
- \fi
-\fi
-\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
-\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
-\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
-\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
-\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
-%%LCP%% \spn@wtheorem{exercise}{Exercise}{\bfseries}{\rmfamily}
-\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
-\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
-\spn@wtheorem{problem}{Problem}{\bfseries}{\rmfamily}
-\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
-\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
-\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
-\spn@wtheorem{solution}{Solution}{\bfseries}{\rmfamily}
-\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
-
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd}
-
-\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
- \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
- \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
- \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
- }
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%%
-%% This is figure.neu
-%%
-%% It redefines the captions for "figure" and "table"
-%% environments.
-%%
-%% There are three new kind of captions: "\firstcaption"
-%% and "\secondcaption" for captions set side by side.
-%% Usage for those two commands: like "\caption".
-%%
-%% "\sidecaption" with two parms: #1 width of picture
-%% #2 height of picture
-%%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\@ifundefined{floatlegendstyle}{\def\floatlegendstyle{\bfseries}}{}
-\def\floatcounterend{.\ }
-\def\capstrut{\vrule\@width\z@\@height\topskip}
-\@ifundefined{captionstyle}{\def\captionstyle{\normalfont\small}}{}
-\@ifundefined{instindent}{\newdimen\instindent}{}
-
-\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-\def\firstcaption{\refstepcounter\@captype\@dblarg%
- {\@firstcaption\@captype}}
-
-\def\secondcaption{\refstepcounter\@captype\@dblarg%
- {\@secondcaption\@captype}}
-
-\long\def\@firstcaption#1[#2]#3{\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \vskip3pt
- \@maketwocaptions{\csname fnum@#1\endcsname}{\ignorespaces #3}%
- \ignorespaces\hspace{.073\textwidth}\hfil%
- \endgroup}
-
-\long\def\@secondcaption#1[#2]#3{\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@maketwocaptions{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-
-\long\def\@maketwocaptions#1#2{%
- \parbox[t]{.46\textwidth}{{\floatlegendstyle #1\floatcounterend} #2}}
-
-\newdimen\figgap\figgap=14.2pt
-%
-\long\def\@makesidecaption#1#2{%
- \setbox0=\vbox{\hsize=\@tempdima
- \captionstyle{\floatlegendstyle
- #1\floatcounterend}#2}%
- \ifdim\instindent<\z@
- \ifdim\ht0>-\instindent
- \advance\instindent by\ht0
- \typeout{^^JClass-Warning: Legend of \string\sidecaption\space for
- \@captype\space\csname the\@captype\endcsname
- ^^Jis \the\instindent\space taller than the corresponding float -
- ^^Jyou'd better switch the environment. }%
- \instindent\z@
- \fi
- \else
- \ifdim\ht0<\instindent
- \advance\instindent by-\ht0
- \advance\instindent by-\dp0\relax
- \advance\instindent by\topskip
- \advance\instindent by-11pt
- \else
- \advance\instindent by-\ht0
- \instindent=-\instindent
- \typeout{^^JClass-Warning: Legend of \string\sidecaption\space for
- \@captype\space\csname the\@captype\endcsname
- ^^Jis \the\instindent\space taller than the corresponding float -
- ^^Jyou'd better switch the environment. }%
- \instindent\z@
- \fi
- \fi
- \parbox[b]{\@tempdima}{\captionstyle{\floatlegendstyle
- #1\floatcounterend}#2%
- \ifdim\instindent>\z@ \\
- \vrule\@width\z@\@height\instindent
- \@depth\z@
- \fi}}
-\def\sidecaption{\@ifnextchar[\sidec@ption{\sidec@ption[b]}}
-\def\sidec@ption[#1]#2\caption{%
-\setbox\@tempboxa=\hbox{\ignorespaces#2\unskip}%
-\if@twocolumn
- \ifdim\hsize<\textwidth\else
- \ifdim\wd\@tempboxa<\columnwidth
- \typeout{Double column float fits into single column -
- ^^Jyou'd better switch the environment. }%
- \fi
- \fi
-\fi
- \instindent=\ht\@tempboxa
- \advance\instindent by\dp\@tempboxa
-\if t#1
-\else
- \instindent=-\instindent
-\fi
-\@tempdima=\hsize
-\advance\@tempdima by-\figgap
-\advance\@tempdima by-\wd\@tempboxa
-\ifdim\@tempdima<3cm
- \typeout{\string\sidecaption: No sufficient room for the legend;
- using normal \string\caption. }%
- \unhbox\@tempboxa
- \let\@capcommand=\@caption
-\else
- \ifdim\@tempdima<4.5cm
- \typeout{\string\sidecaption: Room for the legend very narrow;
- using \string\raggedright. }%
- \toks@\expandafter{\captionstyle\sloppy
- \rightskip=0ptplus6mm\relax}%
- \def\captionstyle{\the\toks@}%
- \fi
- \let\@capcommand=\@sidecaption
- \leavevmode
- \unhbox\@tempboxa
- \hfill
-\fi
-\refstepcounter\@captype
-\@dblarg{\@capcommand\@captype}}
-\long\def\@sidecaption#1[#2]#3{\addcontentsline{\csname
- ext@#1\endcsname}{#1}{\protect\numberline{\csname
- the#1\endcsname}{\ignorespaces #2}}\begingroup
- \@parboxrestore
- \@makesidecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
- \endgroup}
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\fig@type{figure}
-
-\def\leftlegendglue{\hfil}
-\newdimen\figcapgap\figcapgap=3pt
-\newdimen\tabcapgap\tabcapgap=5.5pt
-
-\long\def\@makecaption#1#2{%
- \captionstyle
- \ifx\@captype\fig@type
- \vskip\figcapgap
- \fi
- \setbox\@tempboxa\hbox{{\floatlegendstyle #1\floatcounterend}%
- \capstrut #2}%
- \ifdim \wd\@tempboxa >\hsize
- {\floatlegendstyle #1\floatcounterend}\capstrut #2\par
- \else
- \hbox to\hsize{\leftlegendglue\unhbox\@tempboxa\hfil}%
- \fi
- \ifx\@captype\fig@type\else
- \vskip\tabcapgap
- \fi}
-
-\newcounter{merk}
-\def\endfigure{\resetsubfig\end@float}
-\@namedef{endfigure*}{\resetsubfig\end@dblfloat}
-\let\resetsubfig\relax
-\def\subfigures{\stepcounter{figure}\setcounter{merk}{\value{figure}}%
-\setcounter{figure}{0}\def\thefigure{\if@numart\else\thechapter.\fi
-\@arabic\c@merk\alph{figure}}%
-\def\resetsubfig{\setcounter{figure}{\value{merk}}}}
-\let\leftlegendglue\relax
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% Definition of environment thebibliography
-%
-% Borrowed from book.cls
-%
-% by lcp
-
-\newcommand\bibname{Bibliography}
-\setlength\bibindent{1.5em}
-\renewenvironment{thebibliography}[1]
- {\chapter*{\bibname
- \@mkboth{\MakeUppercase\bibname}{\MakeUppercase\bibname}}%
- \list{\@biblabel{\@arabic\c@enumiv}}%
- {\settowidth\labelwidth{\@biblabel{#1}}%
- \leftmargin\labelwidth
- \advance\leftmargin\labelsep
- \@openbib@code
- \usecounter{enumiv}%
- \let\p@enumiv\@empty
- \renewcommand\theenumiv{\@arabic\c@enumiv}}%
- \sloppy
- \clubpenalty4000
- \@clubpenalty \clubpenalty
- \widowpenalty4000%
- \sfcode`\.\@m}
- {\def\@noitemerr
- {\@latex@warning{Empty `thebibliography' environment}}%
- \endlist}
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is fonotebk.xxx
-%
-% It redefines how footnotes will be typeset.
-%
-% Usage like described by Lamport.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\newdimen\footnoterulewidth
- \footnoterulewidth=1.666cm
-
-\def\footnoterule{\kern-3\p@
- \hrule width\footnoterulewidth
- \kern 2.6\p@}
-
-\newdimen\foot@parindent
-\foot@parindent 10.83\p@
-
-%\long\def\@makefntext#1{\parindent\foot@parindent\noindent
-% \hbox to\foot@parindent{\hss$\m@th^{\@thefnmark}$\kern3pt}#1}
-\long\def\@makefntext#1{\@setpar{\@@par\@tempdima \hsize
- \advance\@tempdima-\foot@parindent\parshape\@ne\foot@parindent
- \@tempdima}\par
- \parindent \foot@parindent\noindent \hbox to \z@{%
- \hss\hss$^{\@thefnmark}$ }#1}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is environ.tex
-%
-% It defines the environment for acknowledgements.
-% and noteadd
-%
-% Usage e.g.: \begin{acknowledgement}
-% Text
-% \end{acknowledgement}
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-% Define `abstract' environment
-\def\acknowledgement{\par\addvspace{17pt}\small\rm
-\trivlist\item[\hskip\labelsep
-{\it\ackname}]}
-\def\endacknowledgement{\endtrivlist\addvspace{6pt}}
-% Define `noteadd' environment
-\def\noteadd{\par\addvspace{17pt}\small\rm
-\trivlist\item[\hskip\labelsep
-{\it\noteaddname}]}
-\def\endnoteadd{\endtrivlist\addvspace{6pt}}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is item.xxx
-%
-% It redefines the kind of label for "itemize", "enumerate"
-% and "description" environment. The last is extended by
-% an optional parameter. Its length is used for overall
-% indentation.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-% labels of enumerate
-
-\def\labelenumi{\theenumi.}
-\def\labelenumii{\theenumii)}
-\def\theenumii{\alph{enumii}}
-\def\p@enumii{\theenumi}
-
-% labels of itemize
-
-\def\labelitemi{\bf --}
-\def\labelitemii{\bf --}
-\def\labelitemiii{$\bullet$}
-\def\labelitemiv{$\cdot$}
-
-% labels of description
-\def\descriptionlabel#1{\hspace\labelsep #1\hfil}
-
-% make indentations changeable
-
-\def\setitemindent#1{\settowidth{\labelwidth}{#1}%
- \leftmargini\labelwidth
- \advance\leftmargini\labelsep
- \def\@listi{\leftmargin\leftmargini
- \labelwidth\leftmargini\advance\labelwidth by -\labelsep
- \parsep=\parskip
- \topsep=\medskipamount
- \itemsep=\parskip \advance\itemsep by -\parsep}}
-\def\setitemitemindent#1{\settowidth{\labelwidth}{#1}%
- \leftmarginii\labelwidth
- \advance\leftmarginii\labelsep
-\def\@listii{\leftmargin\leftmarginii
- \labelwidth\leftmarginii\advance\labelwidth by -\labelsep
- \parsep=\parskip
- \topsep=\z@
- \itemsep=\parskip \advance\itemsep by -\parsep}}
-%
-% adjusted environment "description"
-% if an optional parameter (at the first two levels of lists)
-% is present, its width is considered to be the widest mark
-% throughout the current list.
-\def\description{\@ifnextchar[{\@describe}{\list{}{\labelwidth\z@
- \itemindent-\leftmargin \let\makelabel\descriptionlabel}}}
-%
-\def\describelabel#1{#1\hfil}
-\def\@describe[#1]{\relax\ifnum\@listdepth=0
-\setitemindent{#1}\else\ifnum\@listdepth=1
-\setitemitemindent{#1}\fi\fi
-\list{--}{\let\makelabel\describelabel}}
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is index.xxx
-%
-% It defines miscelaneous addons used for the preparation
-% of an index.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-
-\def\theindex{\@restonecoltrue\if@twocolumn\@restonecolfalse\fi
-\columnseprule \z@
-\columnsep 1cc\twocolumn[\@makeschapterhead{\indexname}%
- \csname indexstarthook\endcsname]%
- \@mkboth{\indexname}{\indexname}%
- \thispagestyle{empty}\parindent\z@
- \rightskip0\p@ plus 40\p@
- \parskip\z@ plus .3\p@\relax\let\item\@idxitem
- \def\,{\relax\ifmmode\mskip\thinmuskip
- \else\hskip0.2em\ignorespaces\fi}%
- \small\rm}
-
-\def\idxquad{\hskip 10\p@}% space that divides entry from number
-
-\def\@idxitem{\par\hangindent 10\p@}
-
-\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
- \noindent\hangindent\wd0\box0}% index entry
-
-\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
- \noindent\hangindent\wd0\box0}% order index entry
-
-\def\endtheindex{\if@restonecol\onecolumn\else\clearpage\fi}
-
-\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
-
-
-
-
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-%
-% This is numberbk.xxx
-%
-% It redefines the kind of numeration for figures,
-% tables and equations. With style option "numart" they
-% are numbered with "no.", otherwise with "kapno.no."
-%
-% e.g. \documentstyle[numart]{article} gives a
-% numbering like in article.sty defined.
-%
-%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
-\def\@takefromreset#1#2{%
- \def\@tempa{#1}%
- \let\@tempd\@elt
- \def\@elt##1{%
- \def\@tempb{##1}%
- \ifx\@tempa\@tempb\else
- \@addtoreset{##1}{#2}%
- \fi}%
- \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
- \expandafter\def\csname cl@#2\endcsname{}%
- \@tempc
- \let\@elt\@tempd
-}
-%
-\def\ds@numart{\@numarttrue
- \@takefromreset{figure}{chapter}%
- \@takefromreset{table}{chapter}%
- \@takefromreset{equation}{chapter}%
- \def\thefigure{\@arabic\c@figure}%
- \def\thetable{\@arabic\c@table}%
- \def\theequation{\arabic{equation}}}
-%
-\def\thefigure{\thechapter.\@arabic\c@figure}
-\def\thetable{\thechapter.\@arabic\c@table}
-\def\theequation{\thechapter.\arabic{equation}}
-\@addtoreset{figure}{chapter}
-\@addtoreset{table}{chapter}
-\@addtoreset{equation}{chapter}
-\endinput
-
--- a/doc-src/TutorialI/document/AB.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,462 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{AB}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isamarkupsection{Case Study: A Context Free Grammar%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-\label{sec:CFG}
-\index{grammars!defining inductively|(}%
-Grammars are nothing but shorthands for inductive definitions of nonterminals
-which represent sets of strings. For example, the production
-$A \to B c$ is short for
-\[ w \in B \Longrightarrow wc \in A \]
-This section demonstrates this idea with an example
-due to Hopcroft and Ullman, a grammar for generating all words with an
-equal number of $a$'s and~$b$'s:
-\begin{eqnarray}
-S &\to& \epsilon \mid b A \mid a B \nonumber\\
-A &\to& a S \mid b A A \nonumber\\
-B &\to& b S \mid a B B \nonumber
-\end{eqnarray}
-At the end we say a few words about the relationship between
-the original proof \cite[p.\ts81]{HopcroftUllman} and our formal version.
-
-We start by fixing the alphabet, which consists only of \isa{a}'s
-and~\isa{b}'s:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{datatype}\isamarkupfalse%
-\ alfa\ {\isaliteral{3D}{\isacharequal}}\ a\ {\isaliteral{7C}{\isacharbar}}\ b%
-\begin{isamarkuptext}%
-\noindent
-For convenience we include the following easy lemmas as simplification rules:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}case{\isaliteral{5F}{\isacharunderscore}}tac\ x{\isaliteral{2C}{\isacharcomma}}\ auto{\isaliteral{29}{\isacharparenright}}%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\noindent
-Words over this alphabet are of type \isa{alfa\ list}, and
-the three nonterminals are declared as sets of such words.
-The productions above are recast as a \emph{mutual} inductive
-definition\index{inductive definition!simultaneous}
-of \isa{S}, \isa{A} and~\isa{B}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}\isamarkupfalse%
-\isanewline
-\ \ S\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}alfa\ list\ set{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline
-\ \ A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}alfa\ list\ set{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline
-\ \ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}alfa\ list\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ S{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ b{\isaliteral{23}{\isacharhash}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ S{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ a{\isaliteral{23}{\isacharhash}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ S{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ S\ \ \ \ \ \ \ \ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ a{\isaliteral{23}{\isacharhash}}w\ \ \ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ v{\isaliteral{5C3C696E3E}{\isasymin}}A{\isaliteral{3B}{\isacharsemicolon}}\ w{\isaliteral{5C3C696E3E}{\isasymin}}A\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ b{\isaliteral{23}{\isacharhash}}v{\isaliteral{40}{\isacharat}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ S\ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ b{\isaliteral{23}{\isacharhash}}w\ \ \ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}\ v\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{3B}{\isacharsemicolon}}\ w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ a{\isaliteral{23}{\isacharhash}}v{\isaliteral{40}{\isacharat}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\noindent
-First we show that all words in \isa{S} contain the same number of \isa{a}'s and \isa{b}'s. Since the definition of \isa{S} is by mutual
-induction, so is the proof: we show at the same time that all words in
-\isa{A} contain one more \isa{a} than \isa{b} and all words in \isa{B} contain one more \isa{b} than \isa{a}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ correctness{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ S\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}b{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ \ \ \ \ {\isaliteral{5C3C616E643E}{\isasymand}}\isanewline
-\ \ \ {\isaliteral{28}{\isacharparenleft}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}b{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\isanewline
-\ \ \ {\isaliteral{28}{\isacharparenleft}}w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}b{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\begin{isamarkuptxt}%
-\noindent
-These propositions are expressed with the help of the predefined \isa{filter} function on lists, which has the convenient syntax \isa{{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}xs{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}}, the list of all elements \isa{x} in \isa{xs} such that \isa{P\ x}
-holds. Remember that on lists \isa{size} and \isa{length} are synonymous.
-
-The proof itself is by rule induction and afterwards automatic:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{by}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}rule\ S{\isaliteral{5F}{\isacharunderscore}}A{\isaliteral{5F}{\isacharunderscore}}B{\isaliteral{2E}{\isachardot}}induct{\isaliteral{2C}{\isacharcomma}}\ auto{\isaliteral{29}{\isacharparenright}}%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\noindent
-This may seem surprising at first, and is indeed an indication of the power
-of inductive definitions. But it is also quite straightforward. For example,
-consider the production $A \to b A A$: if $v,w \in A$ and the elements of $A$
-contain one more $a$ than~$b$'s, then $bvw$ must again contain one more $a$
-than~$b$'s.
-
-As usual, the correctness of syntactic descriptions is easy, but completeness
-is hard: does \isa{S} contain \emph{all} words with an equal number of
-\isa{a}'s and \isa{b}'s? It turns out that this proof requires the
-following lemma: every string with two more \isa{a}'s than \isa{b}'s can be cut somewhere such that each half has one more \isa{a} than
-\isa{b}. This is best seen by imagining counting the difference between the
-number of \isa{a}'s and \isa{b}'s starting at the left end of the
-word. We start with 0 and end (at the right end) with 2. Since each move to the
-right increases or decreases the difference by 1, we must have passed through
-1 on our way from 0 to 2. Formally, we appeal to the following discrete
-intermediate value theorem \isa{nat{\isadigit{0}}{\isaliteral{5F}{\isacharunderscore}}intermed{\isaliteral{5F}{\isacharunderscore}}int{\isaliteral{5F}{\isacharunderscore}}val}
-\begin{isabelle}%
-\ \ \ \ \ {\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}i{\isaliteral{3C}{\isacharless}}n{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6261723E}{\isasymbar}}f\ {\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2D}{\isacharminus}}\ f\ i{\isaliteral{5C3C6261723E}{\isasymbar}}\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isadigit{1}}{\isaliteral{3B}{\isacharsemicolon}}\ f\ {\isadigit{0}}\ {\isaliteral{5C3C6C653E}{\isasymle}}\ k{\isaliteral{3B}{\isacharsemicolon}}\ k\ {\isaliteral{5C3C6C653E}{\isasymle}}\ f\ n{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline
-\isaindent{\ \ \ \ \ }{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}i{\isaliteral{5C3C6C653E}{\isasymle}}n{\isaliteral{2E}{\isachardot}}\ f\ i\ {\isaliteral{3D}{\isacharequal}}\ k%
-\end{isabelle}
-where \isa{f} is of type \isa{nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int}, \isa{int} are the integers,
-\isa{{\isaliteral{5C3C6261723E}{\isasymbar}}{\isaliteral{2E}{\isachardot}}{\isaliteral{5C3C6261723E}{\isasymbar}}} is the absolute value function\footnote{See
-Table~\ref{tab:ascii} in the Appendix for the correct \textsc{ascii}
-syntax.}, and \isa{{\isadigit{1}}} is the integer 1 (see \S\ref{sec:numbers}).
-
-First we show that our specific function, the difference between the
-numbers of \isa{a}'s and \isa{b}'s, does indeed only change by 1 in every
-move to the right. At this point we also start generalizing from \isa{a}'s
-and \isa{b}'s to an arbitrary property \isa{P}. Otherwise we would have
-to prove the desired lemma twice, once as stated above and once with the
-roles of \isa{a}'s and \isa{b}'s interchanged.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ step{\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}i\ {\isaliteral{3C}{\isacharless}}\ size\ w{\isaliteral{2E}{\isachardot}}\isanewline
-\ \ {\isaliteral{5C3C6261723E}{\isasymbar}}{\isaliteral{28}{\isacharparenleft}}int{\isaliteral{28}{\isacharparenleft}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ {\isaliteral{28}{\isacharparenleft}}i{\isaliteral{2B}{\isacharplus}}{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ w{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2D}{\isacharminus}}int{\isaliteral{28}{\isacharparenleft}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ {\isaliteral{28}{\isacharparenleft}}i{\isaliteral{2B}{\isacharplus}}{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ w{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \ {\isaliteral{2D}{\isacharminus}}\ {\isaliteral{28}{\isacharparenleft}}int{\isaliteral{28}{\isacharparenleft}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ w{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2D}{\isacharminus}}int{\isaliteral{28}{\isacharparenleft}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ w{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C6261723E}{\isasymbar}}\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\begin{isamarkuptxt}%
-\noindent
-The lemma is a bit hard to read because of the coercion function
-\isa{int\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ int}. It is required because \isa{size} returns
-a natural number, but subtraction on type~\isa{nat} will do the wrong thing.
-Function \isa{take} is predefined and \isa{take\ i\ xs} is the prefix of
-length \isa{i} of \isa{xs}; below we also need \isa{drop\ i\ xs}, which
-is what remains after that prefix has been dropped from \isa{xs}.
-
-The proof is by induction on \isa{w}, with a trivial base case, and a not
-so trivial induction step. Since it is essentially just arithmetic, we do not
-discuss it.%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ w{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}auto\ simp\ add{\isaliteral{3A}{\isacharcolon}}\ abs{\isaliteral{5F}{\isacharunderscore}}if\ take{\isaliteral{5F}{\isacharunderscore}}Cons\ split{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{2E}{\isachardot}}split{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{done}\isamarkupfalse%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-Finally we come to the above-mentioned lemma about cutting in half a word with two more elements of one sort than of the other sort:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ part{\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ {\isaliteral{22}{\isachardoublequoteopen}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2B}{\isacharplus}}{\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline
-\ \ {\isaliteral{5C3C6578697374733E}{\isasymexists}}i{\isaliteral{5C3C6C653E}{\isasymle}}size\ w{\isaliteral{2E}{\isachardot}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ w{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ w{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2B}{\isacharplus}}{\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\begin{isamarkuptxt}%
-\noindent
-This is proved by \isa{force} with the help of the intermediate value theorem,
-instantiated appropriately and with its first premise disposed of by lemma
-\isa{step{\isadigit{1}}}:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}insert\ nat{\isadigit{0}}{\isaliteral{5F}{\isacharunderscore}}intermed{\isaliteral{5F}{\isacharunderscore}}int{\isaliteral{5F}{\isacharunderscore}}val{\isaliteral{5B}{\isacharbrackleft}}OF\ step{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ of\ {\isaliteral{22}{\isachardoublequoteopen}}P{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}w{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{by}\isamarkupfalse%
-\ force%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\noindent
-
-Lemma \isa{part{\isadigit{1}}} tells us only about the prefix \isa{take\ i\ w}.
-An easy lemma deals with the suffix \isa{drop\ i\ w}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ part{\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ w\ {\isaliteral{40}{\isacharat}}\ drop\ i\ w{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ \ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ w\ {\isaliteral{40}{\isacharat}}\ drop\ i\ w{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2B}{\isacharplus}}{\isadigit{2}}{\isaliteral{3B}{\isacharsemicolon}}\isanewline
-\ \ \ \ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ w{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ w{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2B}{\isacharplus}}{\isadigit{1}}{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline
-\ \ \ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}drop\ i\ w{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}drop\ i\ w{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}P\ x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{2B}{\isacharplus}}{\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{by}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}simp\ del{\isaliteral{3A}{\isacharcolon}}\ append{\isaliteral{5F}{\isacharunderscore}}take{\isaliteral{5F}{\isacharunderscore}}drop{\isaliteral{5F}{\isacharunderscore}}id{\isaliteral{29}{\isacharparenright}}%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\noindent
-In the proof we have disabled the normally useful lemma
-\begin{isabelle}
-\isa{take\ n\ xs\ {\isaliteral{40}{\isacharat}}\ drop\ n\ xs\ {\isaliteral{3D}{\isacharequal}}\ xs}
-\rulename{append_take_drop_id}
-\end{isabelle}
-to allow the simplifier to apply the following lemma instead:
-\begin{isabelle}%
-\ \ \ \ \ {\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C696E3E}{\isasymin}}xs{\isaliteral{40}{\isacharat}}ys{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C696E3E}{\isasymin}}xs{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C696E3E}{\isasymin}}ys{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{5D}{\isacharbrackright}}%
-\end{isabelle}
-
-To dispose of trivial cases automatically, the rules of the inductive
-definition are declared simplification rules:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{declare}\isamarkupfalse%
-\ S{\isaliteral{5F}{\isacharunderscore}}A{\isaliteral{5F}{\isacharunderscore}}B{\isaliteral{2E}{\isachardot}}intros{\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}%
-\begin{isamarkuptext}%
-\noindent
-This could have been done earlier but was not necessary so far.
-
-The completeness theorem tells us that if a word has the same number of
-\isa{a}'s and \isa{b}'s, then it is in \isa{S}, and similarly
-for \isa{A} and \isa{B}:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{theorem}\isamarkupfalse%
-\ completeness{\isaliteral{3A}{\isacharcolon}}\isanewline
-\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}b{\isaliteral{5D}{\isacharbrackright}}\ \ \ \ \ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ S{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\isanewline
-\ \ \ {\isaliteral{28}{\isacharparenleft}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}b{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\isanewline
-\ \ \ {\isaliteral{28}{\isacharparenleft}}size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}b{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ size{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}w{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ w\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\begin{isamarkuptxt}%
-\noindent
-The proof is by induction on \isa{w}. Structural induction would fail here
-because, as we can see from the grammar, we need to make bigger steps than
-merely appending a single letter at the front. Hence we induct on the length
-of \isa{w}, using the induction rule \isa{length{\isaliteral{5F}{\isacharunderscore}}induct}:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ w\ rule{\isaliteral{3A}{\isacharcolon}}\ length{\isaliteral{5F}{\isacharunderscore}}induct{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}rename{\isaliteral{5F}{\isacharunderscore}}tac\ w{\isaliteral{29}{\isacharparenright}}%
-\begin{isamarkuptxt}%
-\noindent
-The \isa{rule} parameter tells \isa{induct{\isaliteral{5F}{\isacharunderscore}}tac} explicitly which induction
-rule to use. For details see \S\ref{sec:complete-ind} below.
-In this case the result is that we may assume the lemma already
-holds for all words shorter than \isa{w}. Because the induction step renames
-the induction variable we rename it back to \isa{w}.
-
-The proof continues with a case distinction on \isa{w},
-on whether \isa{w} is empty or not.%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}case{\isaliteral{5F}{\isacharunderscore}}tac\ w{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{29}{\isacharparenright}}%
-\begin{isamarkuptxt}%
-\noindent
-Simplification disposes of the base case and leaves only a conjunction
-of two step cases to be proved:
-if \isa{w\ {\isaliteral{3D}{\isacharequal}}\ a\ {\isaliteral{23}{\isacharhash}}\ v} and \begin{isabelle}%
-\ \ \ \ \ length\ {\isaliteral{28}{\isacharparenleft}}if\ x\ {\isaliteral{3D}{\isacharequal}}\ a\ then\ {\isaliteral{5B}{\isacharbrackleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ v{\isaliteral{5D}{\isacharbrackright}}\ else\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\isaindent{\ \ \ \ \ }length\ {\isaliteral{28}{\isacharparenleft}}if\ x\ {\isaliteral{3D}{\isacharequal}}\ b\ then\ {\isaliteral{5B}{\isacharbrackleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ v{\isaliteral{5D}{\isacharbrackright}}\ else\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{2}}%
-\end{isabelle} then
-\isa{b\ {\isaliteral{23}{\isacharhash}}\ v\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A}, and similarly for \isa{w\ {\isaliteral{3D}{\isacharequal}}\ b\ {\isaliteral{23}{\isacharhash}}\ v}.
-We only consider the first case in detail.
-
-After breaking the conjunction up into two cases, we can apply
-\isa{part{\isadigit{1}}} to the assumption that \isa{w} contains two more \isa{a}'s than \isa{b}'s.%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}rule\ conjI{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}clarify{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}frule\ part{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}of\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}a{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{2C}{\isacharcomma}}\ simplified{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}clarify{\isaliteral{29}{\isacharparenright}}%
-\begin{isamarkuptxt}%
-\noindent
-This yields an index \isa{i\ {\isaliteral{5C3C6C653E}{\isasymle}}\ length\ v} such that
-\begin{isabelle}%
-\ \ \ \ \ length\ {\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ v\ {\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ length\ {\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}take\ i\ v\ {\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}%
-\end{isabelle}
-With the help of \isa{part{\isadigit{2}}} it follows that
-\begin{isabelle}%
-\ \ \ \ \ length\ {\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}drop\ i\ v\ {\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ length\ {\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}drop\ i\ v\ {\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}%
-\end{isabelle}%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}drule\ part{\isadigit{2}}{\isaliteral{5B}{\isacharbrackleft}}of\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}a{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{2C}{\isacharcomma}}\ simplified{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}assumption{\isaliteral{29}{\isacharparenright}}%
-\begin{isamarkuptxt}%
-\noindent
-Now it is time to decompose \isa{v} in the conclusion \isa{b\ {\isaliteral{23}{\isacharhash}}\ v\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A}
-into \isa{take\ i\ v\ {\isaliteral{40}{\isacharat}}\ drop\ i\ v},%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}rule{\isaliteral{5F}{\isacharunderscore}}tac\ n{\isadigit{1}}{\isaliteral{3D}{\isacharequal}}i\ \isakeyword{and}\ t{\isaliteral{3D}{\isacharequal}}v\ \isakeyword{in}\ subst{\isaliteral{5B}{\isacharbrackleft}}OF\ append{\isaliteral{5F}{\isacharunderscore}}take{\isaliteral{5F}{\isacharunderscore}}drop{\isaliteral{5F}{\isacharunderscore}}id{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}%
-\begin{isamarkuptxt}%
-\noindent
-(the variables \isa{n{\isadigit{1}}} and \isa{t} are the result of composing the
-theorems \isa{subst} and \isa{append{\isaliteral{5F}{\isacharunderscore}}take{\isaliteral{5F}{\isacharunderscore}}drop{\isaliteral{5F}{\isacharunderscore}}id})
-after which the appropriate rule of the grammar reduces the goal
-to the two subgoals \isa{take\ i\ v\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A} and \isa{drop\ i\ v\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A}:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}rule\ S{\isaliteral{5F}{\isacharunderscore}}A{\isaliteral{5F}{\isacharunderscore}}B{\isaliteral{2E}{\isachardot}}intros{\isaliteral{29}{\isacharparenright}}%
-\begin{isamarkuptxt}%
-Both subgoals follow from the induction hypothesis because both \isa{take\ i\ v} and \isa{drop\ i\ v} are shorter than \isa{w}:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\ \ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}force\ simp\ add{\isaliteral{3A}{\isacharcolon}}\ min{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}iff{\isaliteral{5F}{\isacharunderscore}}disj{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}force\ split\ add{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{5F}{\isacharunderscore}}diff{\isaliteral{5F}{\isacharunderscore}}split{\isaliteral{29}{\isacharparenright}}%
-\begin{isamarkuptxt}%
-The case \isa{w\ {\isaliteral{3D}{\isacharequal}}\ b\ {\isaliteral{23}{\isacharhash}}\ v} is proved analogously:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}clarify{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}frule\ part{\isadigit{1}}{\isaliteral{5B}{\isacharbrackleft}}of\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}b{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{2C}{\isacharcomma}}\ simplified{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}clarify{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}drule\ part{\isadigit{2}}{\isaliteral{5B}{\isacharbrackleft}}of\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{3D}{\isacharequal}}b{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{2C}{\isacharcomma}}\ simplified{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}assumption{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}rule{\isaliteral{5F}{\isacharunderscore}}tac\ n{\isadigit{1}}{\isaliteral{3D}{\isacharequal}}i\ \isakeyword{and}\ t{\isaliteral{3D}{\isacharequal}}v\ \isakeyword{in}\ subst{\isaliteral{5B}{\isacharbrackleft}}OF\ append{\isaliteral{5F}{\isacharunderscore}}take{\isaliteral{5F}{\isacharunderscore}}drop{\isaliteral{5F}{\isacharunderscore}}id{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}rule\ S{\isaliteral{5F}{\isacharunderscore}}A{\isaliteral{5F}{\isacharunderscore}}B{\isaliteral{2E}{\isachardot}}intros{\isaliteral{29}{\isacharparenright}}\isanewline
-\ \isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}force\ simp\ add{\isaliteral{3A}{\isacharcolon}}\ min{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}iff{\isaliteral{5F}{\isacharunderscore}}disj{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{by}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}force\ simp\ add{\isaliteral{3A}{\isacharcolon}}\ min{\isaliteral{5F}{\isacharunderscore}}less{\isaliteral{5F}{\isacharunderscore}}iff{\isaliteral{5F}{\isacharunderscore}}disj\ split\ add{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{5F}{\isacharunderscore}}diff{\isaliteral{5F}{\isacharunderscore}}split{\isaliteral{29}{\isacharparenright}}%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-We conclude this section with a comparison of our proof with
-Hopcroft\index{Hopcroft, J. E.} and Ullman's\index{Ullman, J. D.}
-\cite[p.\ts81]{HopcroftUllman}.
-For a start, the textbook
-grammar, for no good reason, excludes the empty word, thus complicating
-matters just a little bit: they have 8 instead of our 7 productions.
-
-More importantly, the proof itself is different: rather than
-separating the two directions, they perform one induction on the
-length of a word. This deprives them of the beauty of rule induction,
-and in the easy direction (correctness) their reasoning is more
-detailed than our \isa{auto}. For the hard part (completeness), they
-consider just one of the cases that our \isa{simp{\isaliteral{5F}{\isacharunderscore}}all} disposes of
-automatically. Then they conclude the proof by saying about the
-remaining cases: ``We do this in a manner similar to our method of
-proof for part (1); this part is left to the reader''. But this is
-precisely the part that requires the intermediate value theorem and
-thus is not at all similar to the other cases (which are automatic in
-Isabelle). The authors are at least cavalier about this point and may
-even have overlooked the slight difficulty lurking in the omitted
-cases. Such errors are found in many pen-and-paper proofs when they
-are scrutinized formally.%
-\index{grammars!defining inductively|)}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/document/ABexpr.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,199 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{ABexpr}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\begin{isamarkuptext}%
-\index{datatypes!mutually recursive}%
-Sometimes it is necessary to define two datatypes that depend on each
-other. This is called \textbf{mutual recursion}. As an example consider a
-language of arithmetic and boolean expressions where
-\begin{itemize}
-\item arithmetic expressions contain boolean expressions because there are
- conditional expressions like ``if $m<n$ then $n-m$ else $m-n$'',
- and
-\item boolean expressions contain arithmetic expressions because of
- comparisons like ``$m<n$''.
-\end{itemize}
-In Isabelle this becomes%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{datatype}\isamarkupfalse%
-\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{3D}{\isacharequal}}\ IF\ \ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7C}{\isacharbar}}\ Sum\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7C}{\isacharbar}}\ Diff\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7C}{\isacharbar}}\ Var\ {\isaliteral{27}{\isacharprime}}a\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7C}{\isacharbar}}\ Num\ nat\isanewline
-\isakeyword{and}\ \ \ \ \ \ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{3D}{\isacharequal}}\ Less\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7C}{\isacharbar}}\ And\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7C}{\isacharbar}}\ Neg\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\noindent
-Type \isa{aexp} is similar to \isa{expr} in \S\ref{sec:ExprCompiler},
-except that we have added an \isa{IF} constructor,
-fixed the values to be of type \isa{nat} and declared the two binary
-operations \isa{Sum} and \isa{Diff}. Boolean
-expressions can be arithmetic comparisons, conjunctions and negations.
-The semantics is given by two evaluation functions:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{primrec}\isamarkupfalse%
-\ evala\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline
-\ \ \ \ \ \ \ \ \ evalb\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}evala\ {\isaliteral{28}{\isacharparenleft}}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ {\isaliteral{28}{\isacharparenleft}}if\ evalb\ b\ env\ then\ evala\ a{\isadigit{1}}\ env\ else\ evala\ a{\isadigit{2}}\ env{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}evala\ {\isaliteral{28}{\isacharparenleft}}Sum\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ evala\ a{\isadigit{1}}\ env\ {\isaliteral{2B}{\isacharplus}}\ evala\ a{\isadigit{2}}\ env{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}evala\ {\isaliteral{28}{\isacharparenleft}}Diff\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ evala\ a{\isadigit{1}}\ env\ {\isaliteral{2D}{\isacharminus}}\ evala\ a{\isadigit{2}}\ env{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}evala\ {\isaliteral{28}{\isacharparenleft}}Var\ v{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ env\ v{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}evala\ {\isaliteral{28}{\isacharparenleft}}Num\ n{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ n{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}evalb\ {\isaliteral{28}{\isacharparenleft}}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}evala\ a{\isadigit{1}}\ env\ {\isaliteral{3C}{\isacharless}}\ evala\ a{\isadigit{2}}\ env{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}evalb\ {\isaliteral{28}{\isacharparenleft}}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}evalb\ b{\isadigit{1}}\ env\ {\isaliteral{5C3C616E643E}{\isasymand}}\ evalb\ b{\isadigit{2}}\ env{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}evalb\ {\isaliteral{28}{\isacharparenleft}}Neg\ b{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ evalb\ b\ env{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\noindent
-
-Both take an expression and an environment (a mapping from variables
-\isa{{\isaliteral{27}{\isacharprime}}a} to values \isa{nat}) and return its arithmetic/boolean
-value. Since the datatypes are mutually recursive, so are functions
-that operate on them. Hence they need to be defined in a single
-\isacommand{primrec} section. Notice the \isakeyword{and} separating
-the declarations of \isa{evala} and \isa{evalb}. Their defining
-equations need not be split into two groups;
-the empty line is purely for readability.
-
-In the same fashion we also define two functions that perform substitution:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{primrec}\isamarkupfalse%
-\ substa\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline
-\ \ \ \ \ \ \ \ \ substb\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}substa\ s\ {\isaliteral{28}{\isacharparenleft}}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\isanewline
-\ \ \ IF\ {\isaliteral{28}{\isacharparenleft}}substb\ s\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}substa\ s\ {\isaliteral{28}{\isacharparenleft}}Sum\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Sum\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}substa\ s\ {\isaliteral{28}{\isacharparenleft}}Diff\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Diff\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}substa\ s\ {\isaliteral{28}{\isacharparenleft}}Var\ v{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ s\ v{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}substa\ s\ {\isaliteral{28}{\isacharparenleft}}Num\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Num\ n{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}substb\ s\ {\isaliteral{28}{\isacharparenleft}}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Less\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}substb\ s\ {\isaliteral{28}{\isacharparenleft}}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ And\ {\isaliteral{28}{\isacharparenleft}}substb\ s\ b{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substb\ s\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline
-{\isaliteral{22}{\isachardoublequoteopen}}substb\ s\ {\isaliteral{28}{\isacharparenleft}}Neg\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Neg\ {\isaliteral{28}{\isacharparenleft}}substb\ s\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-\noindent
-Their first argument is a function mapping variables to expressions, the
-substitution. It is applied to all variables in the second argument. As a
-result, the type of variables in the expression may change from \isa{{\isaliteral{27}{\isacharprime}}a}
-to \isa{{\isaliteral{27}{\isacharprime}}b}. Note that there are only arithmetic and no boolean variables.
-
-Now we can prove a fundamental theorem about the interaction between
-evaluation and substitution: applying a substitution $s$ to an expression $a$
-and evaluating the result in an environment $env$ yields the same result as
-evaluation $a$ in the environment that maps every variable $x$ to the value
-of $s(x)$ under $env$. If you try to prove this separately for arithmetic or
-boolean expressions (by induction), you find that you always need the other
-theorem in the induction step. Therefore you need to state and prove both
-theorems simultaneously:%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ evala\ a\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ evala\ {\isaliteral{28}{\isacharparenleft}}s\ x{\isaliteral{29}{\isacharparenright}}\ env{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\isanewline
-\ \ \ \ \ \ \ \ evalb\ {\isaliteral{28}{\isacharparenleft}}substb\ s\ b{\isaliteral{29}{\isacharparenright}}\ env\ {\isaliteral{3D}{\isacharequal}}\ evalb\ b\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ evala\ {\isaliteral{28}{\isacharparenleft}}s\ x{\isaliteral{29}{\isacharparenright}}\ env{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{apply}\isamarkupfalse%
-{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ a\ \isakeyword{and}\ b{\isaliteral{29}{\isacharparenright}}%
-\begin{isamarkuptxt}%
-\noindent The resulting 8 goals (one for each constructor) are proved in one fell swoop:%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-\isacommand{apply}\isamarkupfalse%
-\ simp{\isaliteral{5F}{\isacharunderscore}}all%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-In general, given $n$ mutually recursive datatypes $\tau@1$, \dots, $\tau@n$,
-an inductive proof expects a goal of the form
-\[ P@1(x@1)\ \land \dots \land P@n(x@n) \]
-where each variable $x@i$ is of type $\tau@i$. Induction is started by
-\begin{isabelle}
-\isacommand{apply}\isa{{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac} $x@1$ \isacommand{and} \dots\ \isacommand{and} $x@n$\isa{{\isaliteral{29}{\isacharparenright}}}
-\end{isabelle}
-
-\begin{exercise}
- Define a function \isa{norma} of type \isa{{\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp} that
- replaces \isa{IF}s with complex boolean conditions by nested
- \isa{IF}s; it should eliminate the constructors
- \isa{And} and \isa{Neg}, leaving only \isa{Less}.
- Prove that \isa{norma}
- preserves the value of an expression and that the result of \isa{norma}
- is really normal, i.e.\ no more \isa{And}s and \isa{Neg}s occur in
- it. ({\em Hint:} proceed as in \S\ref{sec:boolex} and read the discussion
- of type annotations following lemma \isa{subst{\isaliteral{5F}{\isacharunderscore}}id} below).
-\end{exercise}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-\end{isabellebody}%
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: "root"
-%%% End:
--- a/doc-src/TutorialI/document/Advanced.tex Tue Aug 28 13:15:15 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,599 +0,0 @@
-%
-\begin{isabellebody}%
-\def\isabellecontext{Advanced}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isadelimML
-%
-\endisadelimML
-%
-\isatagML
-%
-\endisatagML
-{\isafoldML}%
-%
-\isadelimML
-%
-\endisadelimML
-%
-\begin{isamarkuptext}%
-The premises of introduction rules may contain universal quantifiers and
-monotone functions. A universal quantifier lets the rule
-refer to any number of instances of
-the inductively defined set. A monotone function lets the rule refer
-to existing constructions (such as ``list of'') over the inductively defined
-set. The examples below show how to use the additional expressiveness
-and how to reason from the resulting definitions.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Universal Quantifiers in Introduction Rules \label{sec:gterm-datatype}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-\index{ground terms example|(}%
-\index{quantifiers!and inductive definitions|(}%
-As a running example, this section develops the theory of \textbf{ground
-terms}: terms constructed from constant and function
-symbols but not variables. To simplify matters further, we regard a
-constant as a function applied to the null argument list. Let us declare a
-datatype \isa{gterm} for the type of ground terms. It is a type constructor
-whose argument is a type of function symbols.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{datatype}\isamarkupfalse%
-\ {\isaliteral{27}{\isacharprime}}f\ gterm\ {\isaliteral{3D}{\isacharequal}}\ Apply\ {\isaliteral{27}{\isacharprime}}f\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}f\ gterm\ list{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-To try it out, we declare a datatype of some integer operations:
-integer constants, the unary minus operator and the addition
-operator.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{datatype}\isamarkupfalse%
-\ integer{\isaliteral{5F}{\isacharunderscore}}op\ {\isaliteral{3D}{\isacharequal}}\ Number\ int\ {\isaliteral{7C}{\isacharbar}}\ UnaryMinus\ {\isaliteral{7C}{\isacharbar}}\ Plus%
-\begin{isamarkuptext}%
-Now the type \isa{integer{\isaliteral{5F}{\isacharunderscore}}op\ gterm} denotes the ground
-terms built over those symbols.
-
-The type constructor \isa{gterm} can be generalized to a function
-over sets. It returns
-the set of ground terms that can be formed over a set \isa{F} of function symbols. For
-example, we could consider the set of ground terms formed from the finite
-set \isa{{\isaliteral{7B}{\isacharbraceleft}}Number\ {\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ UnaryMinus{\isaliteral{2C}{\isacharcomma}}\ Plus{\isaliteral{7D}{\isacharbraceright}}}.
-
-This concept is inductive. If we have a list \isa{args} of ground terms
-over~\isa{F} and a function symbol \isa{f} in \isa{F}, then we
-can apply \isa{f} to \isa{args} to obtain another ground term.
-The only difficulty is that the argument list may be of any length. Hitherto,
-each rule in an inductive definition referred to the inductively
-defined set a fixed number of times, typically once or twice.
-A universal quantifier in the premise of the introduction rule
-expresses that every element of \isa{args} belongs
-to our inductively defined set: is a ground term
-over~\isa{F}. The function \isa{set} denotes the set of elements in a given
-list.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}\isamarkupfalse%
-\isanewline
-\ \ gterms\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}f\ set\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}f\ gterm\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isakeyword{for}\ F\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}f\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\isanewline
-step{\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{21}{\isacharbang}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}t\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ args{\isaliteral{2E}{\isachardot}}\ t\ {\isaliteral{5C3C696E3E}{\isasymin}}\ gterms\ F{\isaliteral{3B}{\isacharsemicolon}}\ \ f\ {\isaliteral{5C3C696E3E}{\isasymin}}\ F{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}Apply\ f\ args{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ gterms\ F{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-To demonstrate a proof from this definition, let us
-show that the function \isa{gterms}
-is \textbf{monotone}. We shall need this concept shortly.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ gterms{\isaliteral{5F}{\isacharunderscore}}mono{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}F{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}G\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ gterms\ F\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ gterms\ G{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{apply}\isamarkupfalse%
-\ clarify\isanewline
-\isacommand{apply}\isamarkupfalse%
-\ {\isaliteral{28}{\isacharparenleft}}erule\ gterms{\isaliteral{2E}{\isachardot}}induct{\isaliteral{29}{\isacharparenright}}\isanewline
-\isacommand{apply}\isamarkupfalse%
-\ blast\isanewline
-\isacommand{done}\isamarkupfalse%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\begin{isamarkuptxt}%
-Intuitively, this theorem says that
-enlarging the set of function symbols enlarges the set of ground
-terms. The proof is a trivial rule induction.
-First we use the \isa{clarify} method to assume the existence of an element of
-\isa{gterms\ F}. (We could have used \isa{intro\ subsetI}.) We then
-apply rule induction. Here is the resulting subgoal:
-\begin{isabelle}%
-\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\ args\ f{\isaliteral{2E}{\isachardot}}\isanewline
-\isaindent{\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ \ \ \ }{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}F\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ G{\isaliteral{3B}{\isacharsemicolon}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}t{\isaliteral{5C3C696E3E}{\isasymin}}set\ args{\isaliteral{2E}{\isachardot}}\ t\ {\isaliteral{5C3C696E3E}{\isasymin}}\ gterms\ F\ {\isaliteral{5C3C616E643E}{\isasymand}}\ t\ {\isaliteral{5C3C696E3E}{\isasymin}}\ gterms\ G{\isaliteral{3B}{\isacharsemicolon}}\ f\ {\isaliteral{5C3C696E3E}{\isasymin}}\ F{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline
-\isaindent{\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ \ \ \ }{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Apply\ f\ args\ {\isaliteral{5C3C696E3E}{\isasymin}}\ gterms\ G%
-\end{isabelle}
-The assumptions state that \isa{f} belongs
-to~\isa{F}, which is included in~\isa{G}, and that every element of the list \isa{args} is
-a ground term over~\isa{G}. The \isa{blast} method finds this chain of reasoning easily.%
-\end{isamarkuptxt}%
-\isamarkuptrue%
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\begin{isamarkuptext}%
-\begin{warn}
-Why do we call this function \isa{gterms} instead
-of \isa{gterm}? A constant may have the same name as a type. However,
-name clashes could arise in the theorems that Isabelle generates.
-Our choice of names keeps \isa{gterms{\isaliteral{2E}{\isachardot}}induct} separate from
-\isa{gterm{\isaliteral{2E}{\isachardot}}induct}.
-\end{warn}
-
-Call a term \textbf{well-formed} if each symbol occurring in it is applied
-to the correct number of arguments. (This number is called the symbol's
-\textbf{arity}.) We can express well-formedness by
-generalizing the inductive definition of
-\isa{gterms}.
-Suppose we are given a function called \isa{arity}, specifying the arities
-of all symbols. In the inductive step, we have a list \isa{args} of such
-terms and a function symbol~\isa{f}. If the length of the list matches the
-function's arity then applying \isa{f} to \isa{args} yields a well-formed
-term.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}\isamarkupfalse%
-\isanewline
-\ \ well{\isaliteral{5F}{\isacharunderscore}}formed{\isaliteral{5F}{\isacharunderscore}}gterm\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}f\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}f\ gterm\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\ \ \isakeyword{for}\ arity\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}f\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline
-\isakeyword{where}\isanewline
-step{\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{21}{\isacharbang}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}t\ {\isaliteral{5C3C696E3E}{\isasymin}}\ set\ args{\isaliteral{2E}{\isachardot}}\ t\ {\isaliteral{5C3C696E3E}{\isasymin}}\ well{\isaliteral{5F}{\isacharunderscore}}formed{\isaliteral{5F}{\isacharunderscore}}gterm\ arity{\isaliteral{3B}{\isacharsemicolon}}\ \ \isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ length\ args\ {\isaliteral{3D}{\isacharequal}}\ arity\ f{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}Apply\ f\ args{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ well{\isaliteral{5F}{\isacharunderscore}}formed{\isaliteral{5F}{\isacharunderscore}}gterm\ arity{\isaliteral{22}{\isachardoublequoteclose}}%
-\begin{isamarkuptext}%
-The inductive definition neatly captures the reasoning above.
-The universal quantification over the
-\isa{set} of arguments expresses that all of them are well-formed.%
-\index{quantifiers!and inductive definitions|)}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Alternative Definition Using a Monotone Function%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-\index{monotone functions!and inductive definitions|(}%
-An inductive definition may refer to the
-inductively defined set through an arbitrary monotone function. To
-demonstrate this powerful feature, let us
-change the inductive definition above, replacing the
-quantifier by a use of the function \isa{lists}. This
-function, from the Isabelle theory of lists, is analogous to the
-function \isa{gterms} declared above: if \isa{A} is a set then
-\isa{lists\ A} is the set of lists whose elements belong to
-\isa{A}.
-
-In the inductive definition of well-formed terms, examine the one
-introduction rule. The first premise states that \isa{args} belongs to
-the \isa{lists} of well-formed terms. This formulation is more
-direct, if more obscure, than using a universal quantifier.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}\isamarkupfalse%
-\isanewline
-\ \ well{\isaliteral{5F}{\isacharunderscore}}formed{\isaliteral{5F}{\isacharunderscore}}gterm{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}f\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}f\ gterm\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewl