# HG changeset patch # User kleing # Date 1052732035 -7200 # Node ID 454a2ad0c381b0c8bd96ba5c36a0c01c02fe610b # Parent 75a399c2781f3bed084cf9110fcb5590e357c605 IsarOverview moved one level up diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/IsaMakefile --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/IsaMakefile Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,26 @@ +## targets + +default: Isar + +## global settings + +SRC = $(ISABELLE_HOME)/src +OUT = $(ISABELLE_OUTPUT) +LOG = $(OUT)/log +USEDIR = $(ISATOOL) usedir -i true -d ps -D document -v true + + +## Isar + +Isar: $(LOG)/HOL-Isar.gz + +$(LOG)/HOL-Isar.gz: Isar/ROOT.ML Isar/document/intro.tex \ + Isar/document/root.tex Isar/document/root.bib Isar/*.thy + @$(USEDIR) HOL Isar + + +## clean + +clean: + @rm -f $(LOG)/HOL-Isar.gz + diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/Calc.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/Calc.thy Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,49 @@ +theory Calc = Main: + +subsection{* Chains of equalities *} + +axclass + group < zero, plus, minus + assoc: "(x + y) + z = x + (y + z)" + left_0: "0 + x = x" + left_minus: "-x + x = 0" + +theorem right_minus: "x + -x = (0::'a::group)" +proof - + have "x + -x = (-(-x) + -x) + (x + -x)" + by (simp only: left_0 left_minus) + also have "... = -(-x) + ((-x + x) + -x)" + by (simp only: group.assoc) + also have "... = 0" + by (simp only: left_0 left_minus) + finally show ?thesis . +qed + +subsection{* Inequalities and substitution *} + +lemmas distrib = zadd_zmult_distrib zadd_zmult_distrib2 + zdiff_zmult_distrib zdiff_zmult_distrib2 +declare numeral_2_eq_2[simp] + + +lemma fixes a :: int shows "(a+b)\ \ 2*(a\ + b\)" +proof - + have "(a+b)\ \ (a+b)\ + (a-b)\" by simp + also have "(a+b)\ \ a\ + b\ + 2*a*b" by(simp add:distrib) + also have "(a-b)\ = a\ + b\ - 2*a*b" by(simp add:distrib) + finally show ?thesis by simp +qed + + +subsection{* More transitivity *} + +lemma assumes R: "(a,b) \ R" "(b,c) \ R" "(c,d) \ R" + shows "(a,d) \ R\<^sup>*" +proof - + have "(a,b) \ R\<^sup>*" .. + also have "(b,c) \ R\<^sup>*" .. + also have "(c,d) \ R\<^sup>*" .. + finally show ?thesis . +qed + +end \ No newline at end of file diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/Induction.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/Induction.thy Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,293 @@ +(*<*)theory Induction = Main:(*>*) + +section{*Case distinction and induction \label{sec:Induct}*} + +text{* Computer science applications abound with inductively defined +structures, which is why we treat them in more detail. HOL already +comes with a datatype of lists with the two constructors @{text Nil} +and @{text Cons}. @{text Nil} is written @{term"[]"} and @{text"Cons x +xs"} is written @{term"x # xs"}. *} + +subsection{*Case distinction\label{sec:CaseDistinction}*} + +text{* We have already met the @{text cases} method for performing +binary case splits. Here is another example: *} +lemma "\ A \ A" +proof cases + assume "A" thus ?thesis .. +next + assume "\ A" thus ?thesis .. +qed + +text{*\noindent The two cases must come in this order because @{text +cases} merely abbreviates @{text"(rule case_split_thm)"} where +@{thm[source] case_split_thm} is @{thm case_split_thm}. If we reverse +the order of the two cases in the proof, the first case would prove +@{prop"\ A \ \ A \ A"} which would solve the first premise of +@{thm[source] case_split_thm}, instantiating @{text ?P} with @{term "\ +A"}, thus making the second premise @{prop"\ \ A \ \ A \ A"}. +Therefore the order of subgoals is not always completely arbitrary. + +The above proof is appropriate if @{term A} is textually small. +However, if @{term A} is large, we do not want to repeat it. This can +be avoided by the following idiom *} + +lemma "\ A \ A" +proof (cases "A") + case True thus ?thesis .. +next + case False thus ?thesis .. +qed + +text{*\noindent which is like the previous proof but instantiates +@{text ?P} right away with @{term A}. Thus we could prove the two +cases in any order. The phrase `\isakeyword{case}~@{text True}' +abbreviates `\isakeyword{assume}~@{text"True: A"}' and analogously for +@{text"False"} and @{prop"\A"}. + +The same game can be played with other datatypes, for example lists, +where @{term tl} is the tail of a list, and @{text length} returns a +natural number (remember: $0-1=0$): +*} +(*<*)declare length_tl[simp del](*>*) +lemma "length(tl xs) = length xs - 1" +proof (cases xs) + case Nil thus ?thesis by simp +next + case Cons thus ?thesis by simp +qed +text{*\noindent Here `\isakeyword{case}~@{text Nil}' abbreviates +`\isakeyword{assume}~@{text"Nil:"}~@{prop"xs = []"}' and +`\isakeyword{case}~@{text Cons}' +abbreviates `\isakeyword{fix}~@{text"? ??"} +\isakeyword{assume}~@{text"Cons:"}~@{text"xs = ? # ??"}' +where @{text"?"} and @{text"??"} +stand for variable names that have been chosen by the system. +Therefore we cannot refer to them. +Luckily, this proof is simple enough we do not need to refer to them. +However, sometimes one may have to. Hence Isar offers a simple scheme for +naming those variables: replace the anonymous @{text Cons} by +@{text"(Cons y ys)"}, which abbreviates `\isakeyword{fix}~@{text"y ys"} +\isakeyword{assume}~@{text"Cons:"}~@{text"xs = y # ys"}'. +In each \isakeyword{case} the assumption can be +referred to inside the proof by the name of the constructor. In +Section~\ref{sec:full-Ind} below we will come across an example +of this. *} + +subsection{*Structural induction*} + +text{* We start with an inductive proof where both cases are proved automatically: *} +lemma "2 * (\iii"} or @{text"\"}\label{sec:full-Ind}*} + +text{* Let us now consider the situation where the goal to be proved contains +@{text"\"} or @{text"\"}, say @{prop"\x. P x \ Q x"} --- motivation and a +real example follow shortly. This means that in each case of the induction, +@{text ?case} would be of the form @{prop"\x. P' x \ Q' x"}. Thus the +first proof steps will be the canonical ones, fixing @{text x} and assuming +@{prop"P' x"}. To avoid this tedium, induction performs these steps +automatically: for example in case @{text"(Suc n)"}, @{text ?case} is only +@{prop"Q' x"} whereas the assumptions (named @{term Suc}!) contain both the +usual induction hypothesis \emph{and} @{prop"P' x"}. +It should be clear how this generalises to more complex formulae. + +As an example we will now prove complete induction via +structural induction. *} + +lemma assumes A: "(\n. (\m. m < n \ P m) \ P n)" + shows "P(n::nat)" +proof (rule A) + show "\m. m < n \ P m" + proof (induct n) + case 0 thus ?case by simp + next + case (Suc n) -- {*\isakeyword{fix} @{term m} \isakeyword{assume} @{text Suc}: @{text[source]"?m < n \ P ?m"} @{prop[source]"m < Suc n"}*} + show ?case -- {*@{term ?case}*} + proof cases + assume eq: "m = n" + from Suc and A have "P n" by blast + with eq show "P m" by simp + next + assume "m \ n" + with Suc have "m < n" by arith + thus "P m" by(rule Suc) + qed + qed +qed +text{* \noindent Given the explanations above and the comments in the +proof text (only necessary for novices), the proof should be quite +readable. + +The statement of the lemma is interesting because it deviates from the style in +the Tutorial~\cite{LNCS2283}, which suggests to introduce @{text"\"} or +@{text"\"} into a theorem to strengthen it for induction. In Isar +proofs we can use @{text"\"} and @{text"\"} instead. This simplifies the +proof and means we do not have to convert between the two kinds of +connectives. + +Note that in a nested induction over the same data type, the inner +case labels hide the outer ones of the same name. If you want to refer +to the outer ones inside, you need to name them on the outside, e.g.\ +\isakeyword{note}~@{text"outer_IH = Suc"}. *} + +subsection{*Rule induction*} + +text{* HOL also supports inductively defined sets. See \cite{LNCS2283} +for details. As an example we define our own version of the reflexive +transitive closure of a relation --- HOL provides a predefined one as well.*} +consts rtc :: "('a \ 'a)set \ ('a \ 'a)set" ("_*" [1000] 999) +inductive "r*" +intros +refl: "(x,x) \ r*" +step: "\ (x,y) \ r; (y,z) \ r* \ \ (x,z) \ r*" + +text{* \noindent +First the constant is declared as a function on binary +relations (with concrete syntax @{term"r*"} instead of @{text"rtc +r"}), then the defining clauses are given. We will now prove that +@{term"r*"} is indeed transitive: *} + +lemma assumes A: "(x,y) \ r*" shows "(y,z) \ r* \ (x,z) \ r*" +using A +proof induct + case refl thus ?case . +next + case step thus ?case by(blast intro: rtc.step) +qed +text{*\noindent Rule induction is triggered by a fact $(x_1,\dots,x_n) +\in R$ piped into the proof, here \isakeyword{using}~@{text A}. The +proof itself follows the inductive definition very +closely: there is one case for each rule, and it has the same name as +the rule, analogous to structural induction. + +However, this proof is rather terse. Here is a more readable version: +*} + +lemma assumes A: "(x,y) \ r*" and B: "(y,z) \ r*" + shows "(x,z) \ r*" +proof - + from A B show ?thesis + proof induct + fix x assume "(x,z) \ r*" -- {*@{text B}[@{text y} := @{text x}]*} + thus "(x,z) \ r*" . + next + fix x' x y + assume 1: "(x',x) \ r" and + IH: "(y,z) \ r* \ (x,z) \ r*" and + B: "(y,z) \ r*" + from 1 IH[OF B] show "(x',z) \ r*" by(rule rtc.step) + qed +qed +text{*\noindent We start the proof with \isakeyword{from}~@{text"A +B"}. Only @{text A} is ``consumed'' by the induction step. +Since @{text B} is left over we don't just prove @{text +?thesis} but @{text"B \ ?thesis"}, just as in the previous proof. The +base case is trivial. In the assumptions for the induction step we can +see very clearly how things fit together and permit ourselves the +obvious forward step @{text"IH[OF B]"}. + +The notation `\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}' +is also supported for inductive definitions. The \emph{constructor} is (the +name of) the rule and the \emph{vars} fix the free variables in the +rule; the order of the \emph{vars} must correspond to the +\emph{alphabetical order} of the variables as they appear in the rule. +For example, we could start the above detailed proof of the induction +with \isakeyword{case}~\isa{(step x' x y)}. However, we can then only +refer to the assumptions named \isa{step} collectively and not +individually, as the above proof requires. *} + +subsection{*More induction*} + +text{* We close the section by demonstrating how arbitrary induction +rules are applied. As a simple example we have chosen recursion +induction, i.e.\ induction based on a recursive function +definition. However, most of what we show works for induction in +general. + +The example is an unusual definition of rotation: *} + +consts rot :: "'a list \ 'a list" +recdef rot "measure length" --"for the internal termination proof" +"rot [] = []" +"rot [x] = [x]" +"rot (x#y#zs) = y # rot(x#zs)" +text{*\noindent This yields, among other things, the induction rule +@{thm[source]rot.induct}: @{thm[display]rot.induct[no_vars]} +In the following proof we rely on a default naming scheme for cases: they are +called 1, 2, etc, unless they have been named explicitly. The latter happens +only with datatypes and inductively defined sets, but not with recursive +functions. *} + +lemma "xs \ [] \ rot xs = tl xs @ [hd xs]" +proof (induct xs rule: rot.induct) + case 1 thus ?case by simp +next + case 2 show ?case by simp +next + case (3 a b cs) + have "rot (a # b # cs) = b # rot(a # cs)" by simp + also have "\ = b # tl(a # cs) @ [hd(a # cs)]" by(simp add:3) + also have "\ = tl (a # b # cs) @ [hd (a # b # cs)]" by simp + finally show ?case . +qed + +text{*\noindent +The third case is only shown in gory detail (see \cite{BauerW-TPHOLs01} +for how to reason with chains of equations) to demonstrate that the +`\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}' notation also +works for arbitrary induction theorems with numbered cases. The order +of the \emph{vars} corresponds to the order of the +@{text"\"}-quantified variables in each case of the induction +theorem. For induction theorems produced by \isakeyword{recdef} it is +the order in which the variables appear on the left-hand side of the +equation. + +The proof is so simple that it can be condensed to +*} + +(*<*)lemma "xs \ [] \ rot xs = tl xs @ [hd xs]"(*>*) +by (induct xs rule: rot.induct, simp_all) + +(*<*)end(*>*) diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/Logic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/Logic.thy Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,644 @@ +(*<*)theory Logic = Main:(*>*) + +section{*Logic \label{sec:Logic}*} + +subsection{*Propositional logic*} + +subsubsection{*Introduction rules*} + +text{* We start with a really trivial toy proof to introduce the basic +features of structured proofs. *} +lemma "A \ A" +proof (rule impI) + assume a: "A" + show "A" by(rule a) +qed +text{*\noindent +The operational reading: the \isakeyword{assume}-\isakeyword{show} +block proves @{prop"A \ A"} (@{term a} is a degenerate rule (no +assumptions) that proves @{term A} outright), which rule +@{thm[source]impI} (@{thm impI}) turns into the desired @{prop"A \ +A"}. However, this text is much too detailed for comfort. Therefore +Isar implements the following principle: \begin{quote}\em Command +\isakeyword{proof} automatically tries to select an introduction rule +based on the goal and a predefined list of rules. \end{quote} Here +@{thm[source]impI} is applied automatically: *} + +lemma "A \ A" +proof + assume a: A + show A by(rule a) +qed + +text{*\noindent Single-identifier formulae such as @{term A} need not +be enclosed in double quotes. However, we will continue to do so for +uniformity. + +Trivial proofs, in particular those by assumption, should be trivial +to perform. Proof ``.'' does just that (and a bit more). Thus +naming of assumptions is often superfluous: *} +lemma "A \ A" +proof + assume "A" + show "A" . +qed + +text{* To hide proofs by assumption further, \isakeyword{by}@{text"(method)"} +first applies @{text method} and then tries to solve all remaining subgoals +by assumption: *} +lemma "A \ A \ A" +proof + assume "A" + show "A \ A" by(rule conjI) +qed +text{*\noindent Rule @{thm[source]conjI} is of course @{thm conjI}. +A drawback of implicit proofs by assumption is that it +is no longer obvious where an assumption is used. + +Proofs of the form \isakeyword{by}@{text"(rule"}~\emph{name}@{text")"} +can be abbreviated to ``..'' if \emph{name} refers to one of the +predefined introduction rules (or elimination rules, see below): *} + +lemma "A \ A \ A" +proof + assume "A" + show "A \ A" .. +qed +text{*\noindent +This is what happens: first the matching introduction rule @{thm[source]conjI} +is applied (first ``.''), then the two subgoals are solved by assumption +(second ``.''). *} + +subsubsection{*Elimination rules*} + +text{*A typical elimination rule is @{thm[source]conjE}, $\land$-elimination: +@{thm[display,indent=5]conjE} In the following proof it is applied +by hand, after its first (\emph{major}) premise has been eliminated via +@{text"[OF AB]"}: *} +lemma "A \ B \ B \ A" +proof + assume AB: "A \ B" + show "B \ A" + proof (rule conjE[OF AB]) -- {*@{text"conjE[OF AB]"}: @{thm conjE[OF AB]} *} + assume "A" "B" + show ?thesis .. + qed +qed +text{*\noindent Note that the term @{text"?thesis"} always stands for the +``current goal'', i.e.\ the enclosing \isakeyword{show} (or +\isakeyword{have}) statement. + +This is too much proof text. Elimination rules should be selected +automatically based on their major premise, the formula or rather connective +to be eliminated. In Isar they are triggered by facts being fed +\emph{into} a proof. Syntax: +\begin{center} +\isakeyword{from} \emph{fact} \isakeyword{show} \emph{proposition} \emph{proof} +\end{center} +where \emph{fact} stands for the name of a previously proved +proposition, e.g.\ an assumption, an intermediate result or some global +theorem, which may also be modified with @{text OF} etc. +The \emph{fact} is ``piped'' into the \emph{proof}, which can deal with it +how it chooses. If the \emph{proof} starts with a plain \isakeyword{proof}, +an elimination rule (from a predefined list) is applied +whose first premise is solved by the \emph{fact}. Thus the proof above +is equivalent to the following one: *} + +lemma "A \ B \ B \ A" +proof + assume AB: "A \ B" + from AB show "B \ A" + proof + assume "A" "B" + show ?thesis .. + qed +qed + +text{* Now we come to a second important principle: +\begin{quote}\em +Try to arrange the sequence of propositions in a UNIX-like pipe, +such that the proof of each proposition builds on the previous proposition. +\end{quote} +The previous proposition can be referred to via the fact @{text this}. +This greatly reduces the need for explicit naming of propositions: +*} +lemma "A \ B \ B \ A" +proof + assume "A \ B" + from this show "B \ A" + proof + assume "A" "B" + show ?thesis .. + qed +qed + +text{*\noindent Because of the frequency of \isakeyword{from}~@{text +this}, Isar provides two abbreviations: +\begin{center} +\begin{tabular}{r@ {\quad=\quad}l} +\isakeyword{then} & \isakeyword{from} @{text this} \\ +\isakeyword{thus} & \isakeyword{then} \isakeyword{show} +\end{tabular} +\end{center} + +Here is an alternative proof that operates purely by forward reasoning: *} +lemma "A \ B \ B \ A" +proof + assume ab: "A \ B" + from ab have a: "A" .. + from ab have b: "B" .. + from b a show "B \ A" .. +qed + +text{*\noindent It is worth examining this text in detail because it +exhibits a number of new concepts. For a start, it is the first time +we have proved intermediate propositions (\isakeyword{have}) on the +way to the final \isakeyword{show}. This is the norm in nontrivial +proofs where one cannot bridge the gap between the assumptions and the +conclusion in one step. To understand how the proof works we need to +explain more Isar details. + +Method @{text rule} can be given a list of rules, in which case +@{text"(rule"}~\textit{rules}@{text")"} applies the first matching +rule in the list \textit{rules}. Command \isakeyword{from} can be +followed by any number of facts. Given \isakeyword{from}~@{text +f}$_1$~\dots~@{text f}$_n$, the proof step +@{text"(rule"}~\textit{rules}@{text")"} following a \isakeyword{have} +or \isakeyword{show} searches \textit{rules} for a rule whose first +$n$ premises can be proved by @{text f}$_1$~\dots~@{text f}$_n$ in the +given order. Finally one needs to know that ``..'' is short for +@{text"by(rule"}~\textit{elim-rules intro-rules}@{text")"} (or +@{text"by(rule"}~\textit{intro-rules}@{text")"} if there are no facts +fed into the proof), i.e.\ elimination rules are tried before +introduction rules. + +Thus in the above proof both \isakeyword{have}s are proved via +@{thm[source]conjE} triggered by \isakeyword{from}~@{text ab} whereas +in the \isakeyword{show} step no elimination rule is applicable and +the proof succeeds with @{thm[source]conjI}. The latter would fail had +we written \isakeyword{from}~@{text"a b"} instead of +\isakeyword{from}~@{text"b a"}. + +Proofs starting with a plain @{text proof} behave the same because the +latter is short for @{text"proof (rule"}~\textit{elim-rules +intro-rules}@{text")"} (or @{text"proof +(rule"}~\textit{intro-rules}@{text")"} if there are no facts fed into +the proof). *} + +subsection{*More constructs*} + +text{* In the previous proof of @{prop"A \ B \ B \ A"} we needed to feed +more than one fact into a proof step, a frequent situation. Then the +UNIX-pipe model appears to break down and we need to name the different +facts to refer to them. But this can be avoided: +*} +lemma "A \ B \ B \ A" +proof + assume ab: "A \ B" + from ab have "B" .. + moreover + from ab have "A" .. + ultimately show "B \ A" .. +qed +text{*\noindent You can combine any number of facts @{term A1} \dots\ @{term +An} into a sequence by separating their proofs with +\isakeyword{moreover}. After the final fact, \isakeyword{ultimately} stands +for \isakeyword{from}~@{term A1}~\dots~@{term An}. This avoids having to +introduce names for all of the sequence elements. *} + +text{* Although we have only seen a few introduction and elimination rules so +far, Isar's predefined rules include all the usual natural deduction +rules. We conclude our exposition of propositional logic with an extended +example --- which rules are used implicitly where? *} +lemma "\ (A \ B) \ \ A \ \ B" +proof + assume n: "\ (A \ B)" + show "\ A \ \ B" + proof (rule ccontr) + assume nn: "\ (\ A \ \ B)" + have "\ A" + proof + assume "A" + have "\ B" + proof + assume "B" + have "A \ B" .. + with n show False .. + qed + hence "\ A \ \ B" .. + with nn show False .. + qed + hence "\ A \ \ B" .. + with nn show False .. + qed +qed +text{*\noindent +Rule @{thm[source]ccontr} (``classical contradiction'') is +@{thm ccontr[no_vars]}. +Apart from demonstrating the strangeness of classical +arguments by contradiction, this example also introduces two new +abbreviations: +\begin{center} +\begin{tabular}{l@ {\quad=\quad}l} +\isakeyword{hence} & \isakeyword{then} \isakeyword{have} \\ +\isakeyword{with}~\emph{facts} & +\isakeyword{from}~\emph{facts} @{text this} +\end{tabular} +\end{center} +*} + +subsection{*Avoiding duplication*} + +text{* So far our examples have been a bit unnatural: normally we want to +prove rules expressed with @{text"\"}, not @{text"\"}. Here is an example: +*} +lemma "A \ B \ B \ A" +proof + assume "A \ B" thus "B" .. +next + assume "A \ B" thus "A" .. +qed +text{*\noindent The \isakeyword{proof} always works on the conclusion, +@{prop"B \ A"} in our case, thus selecting $\land$-introduction. Hence +we must show @{prop B} and @{prop A}; both are proved by +$\land$-elimination and the proofs are separated by \isakeyword{next}: +\begin{description} +\item[\isakeyword{next}] deals with multiple subgoals. For example, +when showing @{term"A \ B"} we need to show both @{term A} and @{term +B}. Each subgoal is proved separately, in \emph{any} order. The +individual proofs are separated by \isakeyword{next}. \footnote{Each +\isakeyword{show} must prove one of the pending subgoals. If a +\isakeyword{show} matches multiple subgoals, e.g.\ if the subgoals +contain ?-variables, the first one is proved. Thus the order in which +the subgoals are proved can matter --- see +\S\ref{sec:CaseDistinction} for an example.} + +Strictly speaking \isakeyword{next} is only required if the subgoals +are proved in different assumption contexts which need to be +separated, which is not the case above. For clarity we +have employed \isakeyword{next} anyway and will continue to do so. +\end{description} + +This is all very well as long as formulae are small. Let us now look at some +devices to avoid repeating (possibly large) formulae. A very general method +is pattern matching: *} + +lemma "large_A \ large_B \ large_B \ large_A" + (is "?AB \ ?B \ ?A") +proof + assume "?AB" thus "?B" .. +next + assume "?AB" thus "?A" .. +qed +text{*\noindent Any formula may be followed by +@{text"("}\isakeyword{is}~\emph{pattern}@{text")"} which causes the pattern +to be matched against the formula, instantiating the @{text"?"}-variables in +the pattern. Subsequent uses of these variables in other terms causes +them to be replaced by the terms they stand for. + +We can simplify things even more by stating the theorem by means of the +\isakeyword{assumes} and \isakeyword{shows} elements which allow direct +naming of assumptions: *} + +lemma assumes AB: "large_A \ large_B" + shows "large_B \ large_A" (is "?B \ ?A") +proof + from AB show "?B" .. +next + from AB show "?A" .. +qed +text{*\noindent Note the difference between @{text ?AB}, a term, and +@{text AB}, a fact. + +Finally we want to start the proof with $\land$-elimination so we +don't have to perform it twice, as above. Here is a slick way to +achieve this: *} + +lemma assumes AB: "large_A \ large_B" + shows "large_B \ large_A" (is "?B \ ?A") +using AB +proof + assume "?A" "?B" show ?thesis .. +qed +text{*\noindent Command \isakeyword{using} can appear before a proof +and adds further facts to those piped into the proof. Here @{text AB} +is the only such fact and it triggers $\land$-elimination. Another +frequent idiom is as follows: +\begin{center} +\isakeyword{from} \emph{major-facts}~ +\isakeyword{show} \emph{proposition}~ +\isakeyword{using} \emph{minor-facts}~ +\emph{proof} +\end{center} + +Sometimes it is necessary to suppress the implicit application of rules in a +\isakeyword{proof}. For example \isakeyword{show}~@{prop"A \ B"} would +trigger $\lor$-introduction, requiring us to prove @{prop A}. A simple +``@{text"-"}'' prevents this \emph{faux pas}: *} + +lemma assumes AB: "A \ B" shows "B \ A" +proof - + from AB show ?thesis + proof + assume A show ?thesis .. + next + assume B show ?thesis .. + qed +qed + + +subsection{*Predicate calculus*} + +text{* Command \isakeyword{fix} introduces new local variables into a +proof. The pair \isakeyword{fix}-\isakeyword{show} corresponds to @{text"\"} +(the universal quantifier at the +meta-level) just like \isakeyword{assume}-\isakeyword{show} corresponds to +@{text"\"}. Here is a sample proof, annotated with the rules that are +applied implicitly: *} +lemma assumes P: "\x. P x" shows "\x. P(f x)" +proof --"@{thm[source]allI}: @{thm allI}" + fix a + from P show "P(f a)" .. --"@{thm[source]allE}: @{thm allE}" +qed +text{*\noindent Note that in the proof we have chosen to call the bound +variable @{term a} instead of @{term x} merely to show that the choice of +local names is irrelevant. + +Next we look at @{text"\"} which is a bit more tricky. +*} + +lemma assumes Pf: "\x. P(f x)" shows "\y. P y" +proof - + from Pf show ?thesis + proof -- "@{thm[source]exE}: @{thm exE}" + fix x + assume "P(f x)" + show ?thesis .. -- "@{thm[source]exI}: @{thm exI}" + qed +qed +text{*\noindent Explicit $\exists$-elimination as seen above can become +cumbersome in practice. The derived Isar language element +\isakeyword{obtain} provides a more appealing form of generalised +existence reasoning: *} + +lemma assumes Pf: "\x. P(f x)" shows "\y. P y" +proof - + from Pf obtain x where "P(f x)" .. + thus "\y. P y" .. +qed +text{*\noindent Note how the proof text follows the usual mathematical style +of concluding $P(x)$ from $\exists x. P(x)$, while carefully introducing $x$ +as a new local variable. Technically, \isakeyword{obtain} is similar to +\isakeyword{fix} and \isakeyword{assume} together with a soundness proof of +the elimination involved. + +Here is a proof of a well known tautology. +Which rule is used where? *} + +lemma assumes ex: "\x. \y. P x y" shows "\y. \x. P x y" +proof + fix y + from ex obtain x where "\y. P x y" .. + hence "P x y" .. + thus "\x. P x y" .. +qed + +subsection{*Making bigger steps*} + +text{* So far we have confined ourselves to single step proofs. Of course +powerful automatic methods can be used just as well. Here is an example, +Cantor's theorem that there is no surjective function from a set to its +powerset: *} +theorem "\S. S \ range (f :: 'a \ 'a set)" +proof + let ?S = "{x. x \ f x}" + show "?S \ range f" + proof + assume "?S \ range f" + then obtain y where fy: "?S = f y" .. + show False + proof cases + assume "y \ ?S" + with fy show False by blast + next + assume "y \ ?S" + with fy show False by blast + qed + qed +qed +text{*\noindent +For a start, the example demonstrates two new constructs: +\begin{itemize} +\item \isakeyword{let} introduces an abbreviation for a term, in our case +the witness for the claim. +\item Proof by @{text"cases"} starts a proof by cases. Note that it remains +implicit what the two cases are: it is merely expected that the two subproofs +prove @{text"P \ ?thesis"} and @{text"\P \ ?thesis"} (in that order) +for some @{term P}. +\end{itemize} +If you wonder how to \isakeyword{obtain} @{term y}: +via the predefined elimination rule @{thm rangeE[no_vars]}. + +Method @{text blast} is used because the contradiction does not follow easily +by just a single rule. If you find the proof too cryptic for human +consumption, here is a more detailed version; the beginning up to +\isakeyword{obtain} stays unchanged. *} + +theorem "\S. S \ range (f :: 'a \ 'a set)" +proof + let ?S = "{x. x \ f x}" + show "?S \ range f" + proof + assume "?S \ range f" + then obtain y where fy: "?S = f y" .. + show False + proof cases + assume "y \ ?S" + hence "y \ f y" by simp + hence "y \ ?S" by(simp add:fy) + thus False by contradiction + next + assume "y \ ?S" + hence "y \ f y" by simp + hence "y \ ?S" by(simp add:fy) + thus False by contradiction + qed + qed +qed +text{*\noindent Method @{text contradiction} succeeds if both $P$ and +$\neg P$ are among the assumptions and the facts fed into that step, in any order. + +As it happens, Cantor's theorem can be proved automatically by best-first +search. Depth-first search would diverge, but best-first search successfully +navigates through the large search space: +*} + +theorem "\S. S \ range (f :: 'a \ 'a set)" +by best +(* Of course this only works in the context of HOL's carefully +constructed introduction and elimination rules for set theory.*) + +subsection{*Raw proof blocks*} + +text{* Although we have shown how to employ powerful automatic methods like +@{text blast} to achieve bigger proof steps, there may still be the +tendency to use the default introduction and elimination rules to +decompose goals and facts. This can lead to very tedious proofs: +*} +(*<*)ML"set quick_and_dirty"(*>*) +lemma "\x y. A x y \ B x y \ C x y" +proof + fix x show "\y. A x y \ B x y \ C x y" + proof + fix y show "A x y \ B x y \ C x y" + proof + assume "A x y \ B x y" + show "C x y" sorry + qed + qed +qed +text{*\noindent Since we are only interested in the decomposition and not the +actual proof, the latter has been replaced by +\isakeyword{sorry}. Command \isakeyword{sorry} proves anything but is +only allowed in quick and dirty mode, the default interactive mode. It +is very convenient for top down proof development. + +Luckily we can avoid this step by step decomposition very easily: *} + +lemma "\x y. A x y \ B x y \ C x y" +proof - + have "\x y. \ A x y; B x y \ \ C x y" + proof - + fix x y assume "A x y" "B x y" + show "C x y" sorry + qed + thus ?thesis by blast +qed + +text{*\noindent +This can be simplified further by \emph{raw proof blocks}, i.e.\ +proofs enclosed in braces: *} + +lemma "\x y. A x y \ B x y \ C x y" +proof - + { fix x y assume "A x y" "B x y" + have "C x y" sorry } + thus ?thesis by blast +qed + +text{*\noindent The result of the raw proof block is the same theorem +as above, namely @{prop"\x y. \ A x y; B x y \ \ C x y"}. Raw +proof blocks are like ordinary proofs except that they do not prove +some explicitly stated property but that the property emerges directly +out of the \isakeyword{fixe}s, \isakeyword{assume}s and +\isakeyword{have} in the block. Thus they again serve to avoid +duplication. Note that the conclusion of a raw proof block is stated with +\isakeyword{have} rather than \isakeyword{show} because it is not the +conclusion of some pending goal but some independent claim. + +The general idea demonstrated in this subsection is very +important in Isar and distinguishes it from tactic-style proofs: +\begin{quote}\em +Do not manipulate the proof state into a particular form by applying +tactics but state the desired form explicitly and let the tactic verify +that from this form the original goal follows. +\end{quote} +This yields more readable and also more robust proofs. *} + +subsection{*Further refinements*} + +text{* This subsection discusses some further tricks that can make +life easier although they are not essential. *} + +subsubsection{*\isakeyword{and}*} + +text{* Propositions (following \isakeyword{assume} etc) may but need not be +separated by \isakeyword{and}. This is not just for readability +(\isakeyword{from} \isa{A} \isakeyword{and} \isa{B} looks nicer than +\isakeyword{from} \isa{A} \isa{B}) but for structuring lists of propositions +into possibly named blocks. In +\begin{center} +\isakeyword{assume} \isa{A:} $A_1$ $A_2$ \isakeyword{and} \isa{B:} $A_3$ +\isakeyword{and} $A_4$ +\end{center} +label \isa{A} refers to the list of propositions $A_1$ $A_2$ and +label \isa{B} to $A_3$. *} + +subsubsection{*\isakeyword{note}*} +text{* If you want to remember intermediate fact(s) that cannot be +named directly, use \isakeyword{note}. For example the result of raw +proof block can be named by following it with +\isakeyword{note}~@{text"some_name = this"}. As a side effect, +@{text this} is set to the list of facts on the right-hand side. You +can also say @{text"note some_fact"}, which simply sets @{text this}, +i.e.\ recalls @{text"some_fact"}, e.g.\ in a \isakeyword{moreover} sequence. *} + + +subsubsection{*\isakeyword{fixes}*} + +text{* Sometimes it is necessary to decorate a proposition with type +constraints, as in Cantor's theorem above. These type constraints tend +to make the theorem less readable. The situation can be improved a +little by combining the type constraint with an outer @{text"\"}: *} + +theorem "\f :: 'a \ 'a set. \S. S \ range f" +(*<*)oops(*>*) +text{*\noindent However, now @{term f} is bound and we need a +\isakeyword{fix}~\isa{f} in the proof before we can refer to @{term f}. +This is avoided by \isakeyword{fixes}: *} + +theorem fixes f :: "'a \ 'a set" shows "\S. S \ range f" +(*<*)oops(*>*) +text{* \noindent +Even better, \isakeyword{fixes} allows to introduce concrete syntax locally:*} + +lemma comm_mono: + fixes r :: "'a \ 'a \ bool" (infix ">" 60) and + f :: "'a \ 'a \ 'a" (infixl "++" 70) + assumes comm: "\x y::'a. x ++ y = y ++ x" and + mono: "\x y z::'a. x > y \ x ++ z > y ++ z" + shows "x > y \ z ++ x > z ++ y" +by(simp add: comm mono) + +text{*\noindent The concrete syntax is dropped at the end of the proof and the +theorem becomes @{thm[display,margin=60]comm_mono} +\tweakskip *} + +subsubsection{*\isakeyword{obtain}*} + +text{* The \isakeyword{obtain} construct can introduce multiple +witnesses and propositions as in the following proof fragment:*} + +lemma assumes A: "\x y. P x y \ Q x y" shows "R" +proof - + from A obtain x y where P: "P x y" and Q: "Q x y" by blast +(*<*)oops(*>*) +text{* Remember also that one does not even need to start with a formula +containing @{text"\"} as we saw in the proof of Cantor's theorem. +*} + +subsubsection{*Combining proof styles*} + +text{* Finally, whole ``scripts'' (tactic-based proofs in the style of +\cite{LNCS2283}) may appear in the leaves of the proof tree, although this is +best avoided. Here is a contrived example: *} + +lemma "A \ (A \ B) \ B" +proof + assume a: "A" + show "(A \B) \ B" + apply(rule impI) + apply(erule impE) + apply(rule a) + apply assumption + done +qed + +text{*\noindent You may need to resort to this technique if an +automatic step fails to prove the desired proposition. + +When converting a proof from tactic-style into Isar you can proceed +in a top-down manner: parts of the proof can be left in script form +while the outer structure is already expressed in Isar. *} + +(*<*)end(*>*) diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/ROOT.ML Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,2 @@ +use_thy "Logic"; +use_thy "Induction" diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/.cvsignore --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/.cvsignore Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,2 @@ +*.sty +session.tex \ No newline at end of file diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/Induction.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/Induction.tex Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,505 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Induction}% +\isamarkupfalse% +% +\isamarkupsection{Case distinction and induction \label{sec:Induct}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Computer science applications abound with inductively defined +structures, which is why we treat them in more detail. HOL already +comes with a datatype of lists with the two constructors \isa{Nil} +and \isa{Cons}. \isa{Nil} is written \isa{{\isacharbrackleft}{\isacharbrackright}} and \isa{Cons\ x\ xs} is written \isa{x\ {\isacharhash}\ xs}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Case distinction\label{sec:CaseDistinction}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We have already met the \isa{cases} method for performing +binary case splits. Here is another example:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}{\isasymnot}\ A\ {\isasymor}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ cases\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}{\isasymnot}\ A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent The two cases must come in this order because \isa{cases} merely abbreviates \isa{{\isacharparenleft}rule\ case{\isacharunderscore}split{\isacharunderscore}thm{\isacharparenright}} where +\isa{case{\isacharunderscore}split{\isacharunderscore}thm} is \isa{{\isasymlbrakk}{\isacharquery}P\ {\isasymLongrightarrow}\ {\isacharquery}Q{\isacharsemicolon}\ {\isasymnot}\ {\isacharquery}P\ {\isasymLongrightarrow}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}Q}. If we reverse +the order of the two cases in the proof, the first case would prove +\isa{{\isasymnot}\ A\ {\isasymLongrightarrow}\ {\isasymnot}\ A\ {\isasymor}\ A} which would solve the first premise of +\isa{case{\isacharunderscore}split{\isacharunderscore}thm}, instantiating \isa{{\isacharquery}P} with \isa{{\isasymnot}\ A}, thus making the second premise \isa{{\isasymnot}\ {\isasymnot}\ A\ {\isasymLongrightarrow}\ {\isasymnot}\ A\ {\isasymor}\ A}. +Therefore the order of subgoals is not always completely arbitrary. + +The above proof is appropriate if \isa{A} is textually small. +However, if \isa{A} is large, we do not want to repeat it. This can +be avoided by the following idiom% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}{\isasymnot}\ A\ {\isasymor}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}cases\ {\isachardoublequote}A{\isachardoublequote}{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ True\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ False\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent which is like the previous proof but instantiates +\isa{{\isacharquery}P} right away with \isa{A}. Thus we could prove the two +cases in any order. The phrase `\isakeyword{case}~\isa{True}' +abbreviates `\isakeyword{assume}~\isa{True{\isacharcolon}\ A}' and analogously for +\isa{False} and \isa{{\isasymnot}\ A}. + +The same game can be played with other datatypes, for example lists, +where \isa{tl} is the tail of a list, and \isa{length} returns a +natural number (remember: $0-1=0$):% +\end{isamarkuptext}% +\isamarkuptrue% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}length{\isacharparenleft}tl\ xs{\isacharparenright}\ {\isacharequal}\ length\ xs\ {\isacharminus}\ {\isadigit{1}}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}cases\ xs{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ Nil\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ Cons\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Here `\isakeyword{case}~\isa{Nil}' abbreviates +`\isakeyword{assume}~\isa{Nil{\isacharcolon}}~\isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}}' and +`\isakeyword{case}~\isa{Cons}' +abbreviates `\isakeyword{fix}~\isa{{\isacharquery}\ {\isacharquery}{\isacharquery}} +\isakeyword{assume}~\isa{Cons{\isacharcolon}}~\isa{xs\ {\isacharequal}\ {\isacharquery}\ {\isacharhash}\ {\isacharquery}{\isacharquery}}' +where \isa{{\isacharquery}} and \isa{{\isacharquery}{\isacharquery}} +stand for variable names that have been chosen by the system. +Therefore we cannot refer to them. +Luckily, this proof is simple enough we do not need to refer to them. +However, sometimes one may have to. Hence Isar offers a simple scheme for +naming those variables: replace the anonymous \isa{Cons} by +\isa{{\isacharparenleft}Cons\ y\ ys{\isacharparenright}}, which abbreviates `\isakeyword{fix}~\isa{y\ ys} +\isakeyword{assume}~\isa{Cons{\isacharcolon}}~\isa{xs\ {\isacharequal}\ y\ {\isacharhash}\ ys}'. +In each \isakeyword{case} the assumption can be +referred to inside the proof by the name of the constructor. In +Section~\ref{sec:full-Ind} below we will come across an example +of this.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Structural induction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We start with an inductive proof where both cases are proved automatically:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ {\isacharparenleft}{\isasymSum}i{\isacharless}n{\isacharplus}{\isadigit{1}}{\isachardot}\ i{\isacharparenright}\ {\isacharequal}\ n{\isacharasterisk}{\isacharparenleft}n{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}induct\ n{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent If we want to expose more of the structure of the +proof, we can use pattern matching to avoid having to repeat the goal +statement:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ {\isacharparenleft}{\isasymSum}i{\isacharless}n{\isacharplus}{\isadigit{1}}{\isachardot}\ i{\isacharparenright}\ {\isacharequal}\ n{\isacharasterisk}{\isacharparenleft}n{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequote}{\isacharquery}P\ n{\isachardoublequote}{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}induct\ n{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharquery}P\ {\isadigit{0}}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{fix}\ n\ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}{\isacharquery}P\ n{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}{\isacharquery}P{\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent We could refine this further to show more of the equational +proof. Instead we explore the same avenue as for case distinctions: +introducing context via the \isakeyword{case} command:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ {\isacharparenleft}{\isasymSum}i{\isacharless}n{\isacharplus}{\isadigit{1}}{\isachardot}\ i{\isacharparenright}\ {\isacharequal}\ n{\isacharasterisk}{\isacharparenleft}n{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}induct\ n{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ {\isadigit{0}}\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ Suc\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent The implicitly defined \isa{{\isacharquery}case} refers to the +corresponding case to be proved, i.e.\ \isa{{\isacharquery}P\ {\isadigit{0}}} in the first case and +\isa{{\isacharquery}P{\isacharparenleft}Suc\ n{\isacharparenright}} in the second case. Context \isakeyword{case}~\isa{{\isadigit{0}}} is +empty whereas \isakeyword{case}~\isa{Suc} assumes \isa{{\isacharquery}P\ n}. Again we +have the same problem as with case distinctions: we cannot refer to an anonymous \isa{n} +in the induction step because it has not been introduced via \isakeyword{fix} +(in contrast to the previous proof). The solution is the one outlined for +\isa{Cons} above: replace \isa{Suc} by \isa{{\isacharparenleft}Suc\ i{\isacharparenright}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{fixes}\ n{\isacharcolon}{\isacharcolon}nat\ \isakeyword{shows}\ {\isachardoublequote}n\ {\isacharless}\ n{\isacharasterisk}n\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}induct\ n{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ {\isadigit{0}}\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ {\isacharparenleft}Suc\ i{\isacharparenright}\ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}Suc\ i\ {\isacharless}\ Suc\ i\ {\isacharasterisk}\ Suc\ i\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Of course we could again have written +\isakeyword{thus}~\isa{{\isacharquery}case} instead of giving the term explicitly +but we wanted to use \isa{i} somewhere.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Induction formulae involving \isa{{\isasymAnd}} or \isa{{\isasymLongrightarrow}}\label{sec:full-Ind}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Let us now consider the situation where the goal to be proved contains +\isa{{\isasymAnd}} or \isa{{\isasymLongrightarrow}}, say \isa{{\isasymAnd}x{\isachardot}\ P\ x\ {\isasymLongrightarrow}\ Q\ x} --- motivation and a +real example follow shortly. This means that in each case of the induction, +\isa{{\isacharquery}case} would be of the form \isa{{\isasymAnd}x{\isachardot}\ P{\isacharprime}\ x\ {\isasymLongrightarrow}\ Q{\isacharprime}\ x}. Thus the +first proof steps will be the canonical ones, fixing \isa{x} and assuming +\isa{P{\isacharprime}\ x}. To avoid this tedium, induction performs these steps +automatically: for example in case \isa{{\isacharparenleft}Suc\ n{\isacharparenright}}, \isa{{\isacharquery}case} is only +\isa{Q{\isacharprime}\ x} whereas the assumptions (named \isa{Suc}!) contain both the +usual induction hypothesis \emph{and} \isa{P{\isacharprime}\ x}. +It should be clear how this generalises to more complex formulae. + +As an example we will now prove complete induction via +structural induction.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ A{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isasymAnd}n{\isachardot}\ {\isacharparenleft}{\isasymAnd}m{\isachardot}\ m\ {\isacharless}\ n\ {\isasymLongrightarrow}\ P\ m{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequote}P{\isacharparenleft}n{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}rule\ A{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isasymAnd}m{\isachardot}\ m\ {\isacharless}\ n\ {\isasymLongrightarrow}\ P\ m{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}induct\ n{\isacharparenright}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{case}\ {\isadigit{0}}\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\ \ \isamarkupfalse% +\isacommand{next}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{case}\ {\isacharparenleft}Suc\ n{\isacharparenright}\ \ \ % +\isamarkupcmt{\isakeyword{fix} \isa{m} \isakeyword{assume} \isa{Suc}: \isa{{\isachardoublequote}{\isacharquery}m\ {\isacharless}\ n\ {\isasymLongrightarrow}\ P\ {\isacharquery}m{\isachardoublequote}} \isa{{\isachardoublequote}m\ {\isacharless}\ Suc\ n{\isachardoublequote}}% +} +\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}case\ \ \ \ % +\isamarkupcmt{\isa{P\ m}% +} +\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{proof}\ cases\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ eq{\isacharcolon}\ {\isachardoublequote}m\ {\isacharequal}\ n{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{from}\ Suc\ \isakeyword{and}\ A\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}P\ n{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{with}\ eq\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}P\ m{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{next}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}m\ {\isasymnoteq}\ n{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{with}\ Suc\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}m\ {\isacharless}\ n{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ arith\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}P\ m{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}rule\ Suc{\isacharparenright}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Given the explanations above and the comments in the +proof text (only necessary for novices), the proof should be quite +readable. + +The statement of the lemma is interesting because it deviates from the style in +the Tutorial~\cite{LNCS2283}, which suggests to introduce \isa{{\isasymforall}} or +\isa{{\isasymlongrightarrow}} into a theorem to strengthen it for induction. In Isar +proofs we can use \isa{{\isasymAnd}} and \isa{{\isasymLongrightarrow}} instead. This simplifies the +proof and means we do not have to convert between the two kinds of +connectives. + +Note that in a nested induction over the same data type, the inner +case labels hide the outer ones of the same name. If you want to refer +to the outer ones inside, you need to name them on the outside, e.g.\ +\isakeyword{note}~\isa{outer{\isacharunderscore}IH\ {\isacharequal}\ Suc}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Rule induction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +HOL also supports inductively defined sets. See \cite{LNCS2283} +for details. As an example we define our own version of the reflexive +transitive closure of a relation --- HOL provides a predefined one as well.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{consts}\ rtc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set{\isachardoublequote}\ \ \ {\isacharparenleft}{\isachardoublequote}{\isacharunderscore}{\isacharasterisk}{\isachardoublequote}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{inductive}\ {\isachardoublequote}r{\isacharasterisk}{\isachardoublequote}\isanewline +\isakeyword{intros}\isanewline +refl{\isacharcolon}\ \ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}x{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +step{\isacharcolon}\ \ {\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent +First the constant is declared as a function on binary +relations (with concrete syntax \isa{r{\isacharasterisk}} instead of \isa{rtc\ r}), then the defining clauses are given. We will now prove that +\isa{r{\isacharasterisk}} is indeed transitive:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ A{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{using}\ A\isanewline +\isamarkupfalse% +\isacommand{proof}\ induct\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ refl\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ step\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isachardot}step{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Rule induction is triggered by a fact $(x_1,\dots,x_n) +\in R$ piped into the proof, here \isakeyword{using}~\isa{A}. The +proof itself follows the inductive definition very +closely: there is one case for each rule, and it has the same name as +the rule, analogous to structural induction. + +However, this proof is rather terse. Here is a more readable version:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ A{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\ \isakeyword{and}\ B{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ A\ B\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\ induct\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{fix}\ x\ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\ \ % +\isamarkupcmt{\isa{B}[\isa{y} := \isa{x}]% +} +\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{next}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{fix}\ x{\isacharprime}\ x\ y\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isadigit{1}}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharprime}{\isacharcomma}x{\isacharparenright}\ {\isasymin}\ r{\isachardoublequote}\ \isakeyword{and}\isanewline +\ \ \ \ \ \ \ \ \ \ \ IH{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\ \isakeyword{and}\isanewline +\ \ \ \ \ \ \ \ \ \ \ B{\isacharcolon}\ \ {\isachardoublequote}{\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{from}\ {\isadigit{1}}\ IH{\isacharbrackleft}OF\ B{\isacharbrackright}\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}x{\isacharprime}{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}rule\ rtc{\isachardot}step{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent We start the proof with \isakeyword{from}~\isa{A\ B}. Only \isa{A} is ``consumed'' by the induction step. +Since \isa{B} is left over we don't just prove \isa{{\isacharquery}thesis} but \isa{B\ {\isasymLongrightarrow}\ {\isacharquery}thesis}, just as in the previous proof. The +base case is trivial. In the assumptions for the induction step we can +see very clearly how things fit together and permit ourselves the +obvious forward step \isa{IH{\isacharbrackleft}OF\ B{\isacharbrackright}}. + +The notation `\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}' +is also supported for inductive definitions. The \emph{constructor} is (the +name of) the rule and the \emph{vars} fix the free variables in the +rule; the order of the \emph{vars} must correspond to the +\emph{alphabetical order} of the variables as they appear in the rule. +For example, we could start the above detailed proof of the induction +with \isakeyword{case}~\isa{(step x' x y)}. However, we can then only +refer to the assumptions named \isa{step} collectively and not +individually, as the above proof requires.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{More induction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We close the section by demonstrating how arbitrary induction +rules are applied. As a simple example we have chosen recursion +induction, i.e.\ induction based on a recursive function +definition. However, most of what we show works for induction in +general. + +The example is an unusual definition of rotation:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{consts}\ rot\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{recdef}\ rot\ {\isachardoublequote}measure\ length{\isachardoublequote}\ \ % +\isamarkupcmt{for the internal termination proof% +} +\isanewline +{\isachardoublequote}rot\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline +{\isachardoublequote}rot\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}{\isachardoublequote}\isanewline +{\isachardoublequote}rot\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ y\ {\isacharhash}\ rot{\isacharparenleft}x{\isacharhash}zs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent This yields, among other things, the induction rule +\isa{rot{\isachardot}induct}: \begin{isabelle}% +{\isasymlbrakk}P\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ P\ {\isacharbrackleft}x{\isacharbrackright}{\isacharsemicolon}\ {\isasymAnd}x\ y\ zs{\isachardot}\ P\ {\isacharparenleft}x\ {\isacharhash}\ zs{\isacharparenright}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x% +\end{isabelle} +In the following proof we rely on a default naming scheme for cases: they are +called 1, 2, etc, unless they have been named explicitly. The latter happens +only with datatypes and inductively defined sets, but not with recursive +functions.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ rot\ xs\ {\isacharequal}\ tl\ xs\ {\isacharat}\ {\isacharbrackleft}hd\ xs{\isacharbrackright}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}induct\ xs\ rule{\isacharcolon}\ rot{\isachardot}induct{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ {\isadigit{1}}\ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ {\isadigit{2}}\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{case}\ {\isacharparenleft}{\isadigit{3}}\ a\ b\ cs{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}rot\ {\isacharparenleft}a\ {\isacharhash}\ b\ {\isacharhash}\ cs{\isacharparenright}\ {\isacharequal}\ b\ {\isacharhash}\ rot{\isacharparenleft}a\ {\isacharhash}\ cs{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isasymdots}\ {\isacharequal}\ b\ {\isacharhash}\ tl{\isacharparenleft}a\ {\isacharhash}\ cs{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}hd{\isacharparenleft}a\ {\isacharhash}\ cs{\isacharparenright}{\isacharbrackright}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}{\isadigit{3}}{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isasymdots}\ {\isacharequal}\ tl\ {\isacharparenleft}a\ {\isacharhash}\ b\ {\isacharhash}\ cs{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}hd\ {\isacharparenleft}a\ {\isacharhash}\ b\ {\isacharhash}\ cs{\isacharparenright}{\isacharbrackright}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\ \ \isamarkupfalse% +\isacommand{finally}\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}case\ \isamarkupfalse% +\isacommand{{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent +The third case is only shown in gory detail (see \cite{BauerW-TPHOLs01} +for how to reason with chains of equations) to demonstrate that the +`\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}' notation also +works for arbitrary induction theorems with numbered cases. The order +of the \emph{vars} corresponds to the order of the +\isa{{\isasymAnd}}-quantified variables in each case of the induction +theorem. For induction theorems produced by \isakeyword{recdef} it is +the order in which the variables appear on the left-hand side of the +equation. + +The proof is so simple that it can be condensed to% +\end{isamarkuptext}% +\isamarkuptrue% +\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}induct\ xs\ rule{\isacharcolon}\ rot{\isachardot}induct{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline +\isamarkupfalse% +\isamarkupfalse% +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/Logic.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/Logic.tex Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,1098 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Logic}% +\isamarkupfalse% +% +\isamarkupsection{Logic \label{sec:Logic}% +} +\isamarkuptrue% +% +\isamarkupsubsection{Propositional logic% +} +\isamarkuptrue% +% +\isamarkupsubsubsection{Introduction rules% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We start with a really trivial toy proof to introduce the basic +features of structured proofs.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}rule\ impI{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ a{\isacharcolon}\ {\isachardoublequote}A{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}rule\ a{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent +The operational reading: the \isakeyword{assume}-\isakeyword{show} +block proves \isa{A\ {\isasymLongrightarrow}\ A} (\isa{a} is a degenerate rule (no +assumptions) that proves \isa{A} outright), which rule +\isa{impI} (\isa{{\isacharparenleft}{\isacharquery}P\ {\isasymLongrightarrow}\ {\isacharquery}Q{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isasymlongrightarrow}\ {\isacharquery}Q}) turns into the desired \isa{A\ {\isasymlongrightarrow}\ A}. However, this text is much too detailed for comfort. Therefore +Isar implements the following principle: \begin{quote}\em Command +\isakeyword{proof} automatically tries to select an introduction rule +based on the goal and a predefined list of rules. \end{quote} Here +\isa{impI} is applied automatically:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ a{\isacharcolon}\ A\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ A\ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}rule\ a{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Single-identifier formulae such as \isa{A} need not +be enclosed in double quotes. However, we will continue to do so for +uniformity. + +Trivial proofs, in particular those by assumption, should be trivial +to perform. Proof ``.'' does just that (and a bit more). Thus +naming of assumptions is often superfluous:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +To hide proofs by assumption further, \isakeyword{by}\isa{{\isacharparenleft}method{\isacharparenright}} +first applies \isa{method} and then tries to solve all remaining subgoals +by assumption:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A\ {\isasymand}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}A\ {\isasymand}\ A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}rule\ conjI{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Rule \isa{conjI} is of course \isa{{\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isasymand}\ {\isacharquery}Q}. +A drawback of implicit proofs by assumption is that it +is no longer obvious where an assumption is used. + +Proofs of the form \isakeyword{by}\isa{{\isacharparenleft}rule}~\emph{name}\isa{{\isacharparenright}} +can be abbreviated to ``..'' if \emph{name} refers to one of the +predefined introduction rules (or elimination rules, see below):% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A\ {\isasymand}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}A\ {\isasymand}\ A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent +This is what happens: first the matching introduction rule \isa{conjI} +is applied (first ``.''), then the two subgoals are solved by assumption +(second ``.'').% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Elimination rules% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A typical elimination rule is \isa{conjE}, $\land$-elimination: +\begin{isabelle}% +\ \ \ \ \ {\isasymlbrakk}{\isacharquery}P\ {\isasymand}\ {\isacharquery}Q{\isacharsemicolon}\ {\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R% +\end{isabelle} In the following proof it is applied +by hand, after its first (\emph{major}) premise has been eliminated via +\isa{{\isacharbrackleft}OF\ AB{\isacharbrackright}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ AB{\isacharcolon}\ {\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}rule\ conjE{\isacharbrackleft}OF\ AB{\isacharbrackright}{\isacharparenright}\ \ % +\isamarkupcmt{\isa{conjE{\isacharbrackleft}OF\ AB{\isacharbrackright}}: \isa{{\isacharparenleft}{\isasymlbrakk}A{\isacharsemicolon}\ B{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharquery}R}% +} +\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A{\isachardoublequote}\ {\isachardoublequote}B{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Note that the term \isa{{\isacharquery}thesis} always stands for the +``current goal'', i.e.\ the enclosing \isakeyword{show} (or +\isakeyword{have}) statement. + +This is too much proof text. Elimination rules should be selected +automatically based on their major premise, the formula or rather connective +to be eliminated. In Isar they are triggered by facts being fed +\emph{into} a proof. Syntax: +\begin{center} +\isakeyword{from} \emph{fact} \isakeyword{show} \emph{proposition} \emph{proof} +\end{center} +where \emph{fact} stands for the name of a previously proved +proposition, e.g.\ an assumption, an intermediate result or some global +theorem, which may also be modified with \isa{OF} etc. +The \emph{fact} is ``piped'' into the \emph{proof}, which can deal with it +how it chooses. If the \emph{proof} starts with a plain \isakeyword{proof}, +an elimination rule (from a predefined list) is applied +whose first premise is solved by the \emph{fact}. Thus the proof above +is equivalent to the following one:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ AB{\isacharcolon}\ {\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ AB\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A{\isachardoublequote}\ {\isachardoublequote}B{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +Now we come to a second important principle: +\begin{quote}\em +Try to arrange the sequence of propositions in a UNIX-like pipe, +such that the proof of each proposition builds on the previous proposition. +\end{quote} +The previous proposition can be referred to via the fact \isa{this}. +This greatly reduces the need for explicit naming of propositions:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ this\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A{\isachardoublequote}\ {\isachardoublequote}B{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Because of the frequency of \isakeyword{from}~\isa{this}, Isar provides two abbreviations: +\begin{center} +\begin{tabular}{r@ {\quad=\quad}l} +\isakeyword{then} & \isakeyword{from} \isa{this} \\ +\isakeyword{thus} & \isakeyword{then} \isakeyword{show} +\end{tabular} +\end{center} + +Here is an alternative proof that operates purely by forward reasoning:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ ab{\isacharcolon}\ {\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ ab\ \isamarkupfalse% +\isacommand{have}\ a{\isacharcolon}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ ab\ \isamarkupfalse% +\isacommand{have}\ b{\isacharcolon}\ {\isachardoublequote}B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ b\ a\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}B\ {\isasymand}\ A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent It is worth examining this text in detail because it +exhibits a number of new concepts. For a start, it is the first time +we have proved intermediate propositions (\isakeyword{have}) on the +way to the final \isakeyword{show}. This is the norm in nontrivial +proofs where one cannot bridge the gap between the assumptions and the +conclusion in one step. To understand how the proof works we need to +explain more Isar details. + +Method \isa{rule} can be given a list of rules, in which case +\isa{{\isacharparenleft}rule}~\textit{rules}\isa{{\isacharparenright}} applies the first matching +rule in the list \textit{rules}. Command \isakeyword{from} can be +followed by any number of facts. Given \isakeyword{from}~\isa{f}$_1$~\dots~\isa{f}$_n$, the proof step +\isa{{\isacharparenleft}rule}~\textit{rules}\isa{{\isacharparenright}} following a \isakeyword{have} +or \isakeyword{show} searches \textit{rules} for a rule whose first +$n$ premises can be proved by \isa{f}$_1$~\dots~\isa{f}$_n$ in the +given order. Finally one needs to know that ``..'' is short for +\isa{by{\isacharparenleft}rule}~\textit{elim-rules intro-rules}\isa{{\isacharparenright}} (or +\isa{by{\isacharparenleft}rule}~\textit{intro-rules}\isa{{\isacharparenright}} if there are no facts +fed into the proof), i.e.\ elimination rules are tried before +introduction rules. + +Thus in the above proof both \isakeyword{have}s are proved via +\isa{conjE} triggered by \isakeyword{from}~\isa{ab} whereas +in the \isakeyword{show} step no elimination rule is applicable and +the proof succeeds with \isa{conjI}. The latter would fail had +we written \isakeyword{from}~\isa{a\ b} instead of +\isakeyword{from}~\isa{b\ a}. + +Proofs starting with a plain \isa{proof} behave the same because the +latter is short for \isa{proof\ {\isacharparenleft}rule}~\textit{elim-rules +intro-rules}\isa{{\isacharparenright}} (or \isa{proof\ {\isacharparenleft}rule}~\textit{intro-rules}\isa{{\isacharparenright}} if there are no facts fed into +the proof).% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{More constructs% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In the previous proof of \isa{A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A} we needed to feed +more than one fact into a proof step, a frequent situation. Then the +UNIX-pipe model appears to break down and we need to name the different +facts to refer to them. But this can be avoided:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ ab{\isacharcolon}\ {\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ ab\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{moreover}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ ab\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{ultimately}\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}B\ {\isasymand}\ A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent You can combine any number of facts \isa{A{\isadigit{1}}} \dots\ \isa{An} into a sequence by separating their proofs with +\isakeyword{moreover}. After the final fact, \isakeyword{ultimately} stands +for \isakeyword{from}~\isa{A{\isadigit{1}}}~\dots~\isa{An}. This avoids having to +introduce names for all of the sequence elements.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +Although we have only seen a few introduction and elimination rules so +far, Isar's predefined rules include all the usual natural deduction +rules. We conclude our exposition of propositional logic with an extended +example --- which rules are used implicitly where?% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}{\isasymnot}\ {\isacharparenleft}A\ {\isasymand}\ B{\isacharparenright}\ {\isasymlongrightarrow}\ {\isasymnot}\ A\ {\isasymor}\ {\isasymnot}\ B{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ n{\isacharcolon}\ {\isachardoublequote}{\isasymnot}\ {\isacharparenleft}A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isasymnot}\ A\ {\isasymor}\ {\isasymnot}\ B{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}rule\ ccontr{\isacharparenright}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ nn{\isacharcolon}\ {\isachardoublequote}{\isasymnot}\ {\isacharparenleft}{\isasymnot}\ A\ {\isasymor}\ {\isasymnot}\ B{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isasymnot}\ A{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isasymnot}\ B{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}B{\isachardoublequote}\isanewline +\ \ \ \ \ \ \ \ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \ \ \ \ \ \ \isamarkupfalse% +\isacommand{with}\ n\ \isamarkupfalse% +\isacommand{show}\ False\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{hence}\ {\isachardoublequote}{\isasymnot}\ A\ {\isasymor}\ {\isasymnot}\ B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{with}\ nn\ \isamarkupfalse% +\isacommand{show}\ False\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{hence}\ {\isachardoublequote}{\isasymnot}\ A\ {\isasymor}\ {\isasymnot}\ B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{with}\ nn\ \isamarkupfalse% +\isacommand{show}\ False\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent +Rule \isa{ccontr} (``classical contradiction'') is +\isa{{\isacharparenleft}{\isasymnot}\ P\ {\isasymLongrightarrow}\ False{\isacharparenright}\ {\isasymLongrightarrow}\ P}. +Apart from demonstrating the strangeness of classical +arguments by contradiction, this example also introduces two new +abbreviations: +\begin{center} +\begin{tabular}{l@ {\quad=\quad}l} +\isakeyword{hence} & \isakeyword{then} \isakeyword{have} \\ +\isakeyword{with}~\emph{facts} & +\isakeyword{from}~\emph{facts} \isa{this} +\end{tabular} +\end{center}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Avoiding duplication% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +So far our examples have been a bit unnatural: normally we want to +prove rules expressed with \isa{{\isasymLongrightarrow}}, not \isa{{\isasymlongrightarrow}}. Here is an example:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A\ {\isasymand}\ B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent The \isakeyword{proof} always works on the conclusion, +\isa{B\ {\isasymand}\ A} in our case, thus selecting $\land$-introduction. Hence +we must show \isa{B} and \isa{A}; both are proved by +$\land$-elimination and the proofs are separated by \isakeyword{next}: +\begin{description} +\item[\isakeyword{next}] deals with multiple subgoals. For example, +when showing \isa{A\ {\isasymand}\ B} we need to show both \isa{A} and \isa{B}. Each subgoal is proved separately, in \emph{any} order. The +individual proofs are separated by \isakeyword{next}. \footnote{Each +\isakeyword{show} must prove one of the pending subgoals. If a +\isakeyword{show} matches multiple subgoals, e.g.\ if the subgoals +contain ?-variables, the first one is proved. Thus the order in which +the subgoals are proved can matter --- see +\S\ref{sec:CaseDistinction} for an example.} + +Strictly speaking \isakeyword{next} is only required if the subgoals +are proved in different assumption contexts which need to be +separated, which is not the case above. For clarity we +have employed \isakeyword{next} anyway and will continue to do so. +\end{description} + +This is all very well as long as formulae are small. Let us now look at some +devices to avoid repeating (possibly large) formulae. A very general method +is pattern matching:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B\ {\isasymLongrightarrow}\ large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequote}\isanewline +\ \ \ \ \ \ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequote}{\isacharquery}AB\ {\isasymLongrightarrow}\ {\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequote}{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}{\isacharquery}AB{\isachardoublequote}\ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}{\isacharquery}B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}{\isacharquery}AB{\isachardoublequote}\ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}{\isacharquery}A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Any formula may be followed by +\isa{{\isacharparenleft}}\isakeyword{is}~\emph{pattern}\isa{{\isacharparenright}} which causes the pattern +to be matched against the formula, instantiating the \isa{{\isacharquery}}-variables in +the pattern. Subsequent uses of these variables in other terms causes +them to be replaced by the terms they stand for. + +We can simplify things even more by stating the theorem by means of the +\isakeyword{assumes} and \isakeyword{shows} elements which allow direct +naming of assumptions:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequote}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequote}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequote}large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequote}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequote}{\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequote}{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ AB\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharquery}B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{next}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ AB\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharquery}A{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Note the difference between \isa{{\isacharquery}AB}, a term, and +\isa{AB}, a fact. + +Finally we want to start the proof with $\land$-elimination so we +don't have to perform it twice, as above. Here is a slick way to +achieve this:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequote}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequote}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequote}large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequote}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequote}{\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequote}{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{using}\ AB\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}{\isacharquery}A{\isachardoublequote}\ {\isachardoublequote}{\isacharquery}B{\isachardoublequote}\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Command \isakeyword{using} can appear before a proof +and adds further facts to those piped into the proof. Here \isa{AB} +is the only such fact and it triggers $\land$-elimination. Another +frequent idiom is as follows: +\begin{center} +\isakeyword{from} \emph{major-facts}~ +\isakeyword{show} \emph{proposition}~ +\isakeyword{using} \emph{minor-facts}~ +\emph{proof} +\end{center} + +Sometimes it is necessary to suppress the implicit application of rules in a +\isakeyword{proof}. For example \isakeyword{show}~\isa{A\ {\isasymor}\ B} would +trigger $\lor$-introduction, requiring us to prove \isa{A}. A simple +``\isa{{\isacharminus}}'' prevents this \emph{faux pas}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequote}A\ {\isasymor}\ B{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}B\ {\isasymor}\ A{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ AB\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ A\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{next}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ B\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\isamarkupsubsection{Predicate calculus% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Command \isakeyword{fix} introduces new local variables into a +proof. The pair \isakeyword{fix}-\isakeyword{show} corresponds to \isa{{\isasymAnd}} +(the universal quantifier at the +meta-level) just like \isakeyword{assume}-\isakeyword{show} corresponds to +\isa{{\isasymLongrightarrow}}. Here is a sample proof, annotated with the rules that are +applied implicitly:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ P{\isacharcolon}\ {\isachardoublequote}{\isasymforall}x{\isachardot}\ P\ x{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymforall}x{\isachardot}\ P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ % +\isamarkupcmt{\isa{allI}: \isa{{\isacharparenleft}{\isasymAnd}x{\isachardot}\ {\isacharquery}P\ x{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymforall}x{\isachardot}\ {\isacharquery}P\ x}% +} +\isanewline +\ \ \isamarkupfalse% +\isacommand{fix}\ a\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ P\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}P{\isacharparenleft}f\ a{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\ \ % +\isamarkupcmt{\isa{allE}: \isa{{\isasymlbrakk}{\isasymforall}x{\isachardot}\ {\isacharquery}P\ x{\isacharsemicolon}\ {\isacharquery}P\ {\isacharquery}x\ {\isasymLongrightarrow}\ {\isacharquery}R{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}R}% +} +\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Note that in the proof we have chosen to call the bound +variable \isa{a} instead of \isa{x} merely to show that the choice of +local names is irrelevant. + +Next we look at \isa{{\isasymexists}} which is a bit more tricky.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ Pf{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x{\isachardot}\ P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymexists}y{\isachardot}\ P\ y{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ Pf\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\ \ \ \ \ \ \ \ \ \ \ \ \ \ % +\isamarkupcmt{\isa{exE}: \isa{{\isasymlbrakk}{\isasymexists}x{\isachardot}\ {\isacharquery}P\ x{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ {\isacharquery}P\ x\ {\isasymLongrightarrow}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}Q}% +} +\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{fix}\ x\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\ \ % +\isamarkupcmt{\isa{exI}: \isa{{\isacharquery}P\ {\isacharquery}x\ {\isasymLongrightarrow}\ {\isasymexists}x{\isachardot}\ {\isacharquery}P\ x}% +} +\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Explicit $\exists$-elimination as seen above can become +cumbersome in practice. The derived Isar language element +\isakeyword{obtain} provides a more appealing form of generalised +existence reasoning:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ Pf{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x{\isachardot}\ P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymexists}y{\isachardot}\ P\ y{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ Pf\ \isamarkupfalse% +\isacommand{obtain}\ x\ \isakeyword{where}\ {\isachardoublequote}P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}{\isasymexists}y{\isachardot}\ P\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Note how the proof text follows the usual mathematical style +of concluding $P(x)$ from $\exists x. P(x)$, while carefully introducing $x$ +as a new local variable. Technically, \isakeyword{obtain} is similar to +\isakeyword{fix} and \isakeyword{assume} together with a soundness proof of +the elimination involved. + +Here is a proof of a well known tautology. +Which rule is used where?% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ ex{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x{\isachardot}\ {\isasymforall}y{\isachardot}\ P\ x\ y{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymforall}y{\isachardot}\ {\isasymexists}x{\isachardot}\ P\ x\ y{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{fix}\ y\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ ex\ \isamarkupfalse% +\isacommand{obtain}\ x\ \isakeyword{where}\ {\isachardoublequote}{\isasymforall}y{\isachardot}\ P\ x\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{hence}\ {\isachardoublequote}P\ x\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \isamarkupfalse% +\isacommand{thus}\ {\isachardoublequote}{\isasymexists}x{\isachardot}\ P\ x\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\isamarkupsubsection{Making bigger steps% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +So far we have confined ourselves to single step proofs. Of course +powerful automatic methods can be used just as well. Here is an example, +Cantor's theorem that there is no surjective function from a set to its +powerset:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{theorem}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isacharparenright}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{let}\ {\isacharquery}S\ {\isacharequal}\ {\isachardoublequote}{\isacharbraceleft}x{\isachardot}\ x\ {\isasymnotin}\ f\ x{\isacharbraceright}{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharquery}S\ {\isasymnotin}\ range\ f{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}{\isacharquery}S\ {\isasymin}\ range\ f{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{then}\ \isamarkupfalse% +\isacommand{obtain}\ y\ \isakeyword{where}\ fy{\isacharcolon}\ {\isachardoublequote}{\isacharquery}S\ {\isacharequal}\ f\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{show}\ False\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{proof}\ cases\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}y\ {\isasymin}\ {\isacharquery}S{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{with}\ fy\ \isamarkupfalse% +\isacommand{show}\ False\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{next}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}y\ {\isasymnotin}\ {\isacharquery}S{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{with}\ fy\ \isamarkupfalse% +\isacommand{show}\ False\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent +For a start, the example demonstrates two new constructs: +\begin{itemize} +\item \isakeyword{let} introduces an abbreviation for a term, in our case +the witness for the claim. +\item Proof by \isa{cases} starts a proof by cases. Note that it remains +implicit what the two cases are: it is merely expected that the two subproofs +prove \isa{P\ {\isasymLongrightarrow}\ {\isacharquery}thesis} and \isa{{\isasymnot}P\ {\isasymLongrightarrow}\ {\isacharquery}thesis} (in that order) +for some \isa{P}. +\end{itemize} +If you wonder how to \isakeyword{obtain} \isa{y}: +via the predefined elimination rule \isa{{\isasymlbrakk}b\ {\isasymin}\ range\ f{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ b\ {\isacharequal}\ f\ x\ {\isasymLongrightarrow}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ P}. + +Method \isa{blast} is used because the contradiction does not follow easily +by just a single rule. If you find the proof too cryptic for human +consumption, here is a more detailed version; the beginning up to +\isakeyword{obtain} stays unchanged.% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{theorem}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isacharparenright}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{let}\ {\isacharquery}S\ {\isacharequal}\ {\isachardoublequote}{\isacharbraceleft}x{\isachardot}\ x\ {\isasymnotin}\ f\ x{\isacharbraceright}{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharquery}S\ {\isasymnotin}\ range\ f{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}{\isacharquery}S\ {\isasymin}\ range\ f{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{then}\ \isamarkupfalse% +\isacommand{obtain}\ y\ \isakeyword{where}\ fy{\isacharcolon}\ {\isachardoublequote}{\isacharquery}S\ {\isacharequal}\ f\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{show}\ False\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{proof}\ cases\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}y\ {\isasymin}\ {\isacharquery}S{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{hence}\ {\isachardoublequote}y\ {\isasymnotin}\ f\ y{\isachardoublequote}\ \ \ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{hence}\ {\isachardoublequote}y\ {\isasymnotin}\ {\isacharquery}S{\isachardoublequote}\ \ \ \ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}fy{\isacharparenright}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{thus}\ False\ \ \ \ \ \ \ \ \ \isamarkupfalse% +\isacommand{by}\ contradiction\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{next}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}y\ {\isasymnotin}\ {\isacharquery}S{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{hence}\ {\isachardoublequote}y\ {\isasymin}\ f\ y{\isachardoublequote}\ \ \ \isamarkupfalse% +\isacommand{by}\ simp\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{hence}\ {\isachardoublequote}y\ {\isasymin}\ {\isacharquery}S{\isachardoublequote}\ \ \ \ \isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}fy{\isacharparenright}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{thus}\ False\ \ \ \ \ \ \ \ \ \isamarkupfalse% +\isacommand{by}\ contradiction\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Method \isa{contradiction} succeeds if both $P$ and +$\neg P$ are among the assumptions and the facts fed into that step, in any order. + +As it happens, Cantor's theorem can be proved automatically by best-first +search. Depth-first search would diverge, but best-first search successfully +navigates through the large search space:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{theorem}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isacharparenright}{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{by}\ best\isamarkupfalse% +% +\isamarkupsubsection{Raw proof blocks% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Although we have shown how to employ powerful automatic methods like +\isa{blast} to achieve bigger proof steps, there may still be the +tendency to use the default introduction and elimination rules to +decompose goals and facts. This can lead to very tedious proofs:% +\end{isamarkuptext}% +\isamarkuptrue% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x\ y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{fix}\ x\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isasymforall}y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{fix}\ y\ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A\ x\ y\ {\isasymand}\ B\ x\ y{\isachardoublequote}\isanewline +\ \ \ \ \ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}C\ x\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{sorry}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent Since we are only interested in the decomposition and not the +actual proof, the latter has been replaced by +\isakeyword{sorry}. Command \isakeyword{sorry} proves anything but is +only allowed in quick and dirty mode, the default interactive mode. It +is very convenient for top down proof development. + +Luckily we can avoid this step by step decomposition very easily:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x\ y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isasymAnd}x\ y{\isachardot}\ {\isasymlbrakk}\ A\ x\ y{\isacharsemicolon}\ B\ x\ y\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{fix}\ x\ y\ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A\ x\ y{\isachardoublequote}\ {\isachardoublequote}B\ x\ y{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}C\ x\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{sorry}\isanewline +\ \ \isamarkupfalse% +\isacommand{qed}\isanewline +\ \ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent +This can be simplified further by \emph{raw proof blocks}, i.e.\ +proofs enclosed in braces:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x\ y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{{\isacharbraceleft}}\ \isamarkupfalse% +\isacommand{fix}\ x\ y\ \isamarkupfalse% +\isacommand{assume}\ {\isachardoublequote}A\ x\ y{\isachardoublequote}\ {\isachardoublequote}B\ x\ y{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}C\ x\ y{\isachardoublequote}\ \isamarkupfalse% +\isacommand{sorry}\ \isamarkupfalse% +\isacommand{{\isacharbraceright}}\isanewline +\ \ \isamarkupfalse% +\isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent The result of the raw proof block is the same theorem +as above, namely \isa{{\isasymAnd}x\ y{\isachardot}\ {\isasymlbrakk}A\ x\ y{\isacharsemicolon}\ B\ x\ y{\isasymrbrakk}\ {\isasymLongrightarrow}\ C\ x\ y}. Raw +proof blocks are like ordinary proofs except that they do not prove +some explicitly stated property but that the property emerges directly +out of the \isakeyword{fixe}s, \isakeyword{assume}s and +\isakeyword{have} in the block. Thus they again serve to avoid +duplication. Note that the conclusion of a raw proof block is stated with +\isakeyword{have} rather than \isakeyword{show} because it is not the +conclusion of some pending goal but some independent claim. + +The general idea demonstrated in this subsection is very +important in Isar and distinguishes it from tactic-style proofs: +\begin{quote}\em +Do not manipulate the proof state into a particular form by applying +tactics but state the desired form explicitly and let the tactic verify +that from this form the original goal follows. +\end{quote} +This yields more readable and also more robust proofs.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Further refinements% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +This subsection discusses some further tricks that can make +life easier although they are not essential.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{\isakeyword{and}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Propositions (following \isakeyword{assume} etc) may but need not be +separated by \isakeyword{and}. This is not just for readability +(\isakeyword{from} \isa{A} \isakeyword{and} \isa{B} looks nicer than +\isakeyword{from} \isa{A} \isa{B}) but for structuring lists of propositions +into possibly named blocks. In +\begin{center} +\isakeyword{assume} \isa{A:} $A_1$ $A_2$ \isakeyword{and} \isa{B:} $A_3$ +\isakeyword{and} $A_4$ +\end{center} +label \isa{A} refers to the list of propositions $A_1$ $A_2$ and +label \isa{B} to $A_3$.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{\isakeyword{note}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +If you want to remember intermediate fact(s) that cannot be +named directly, use \isakeyword{note}. For example the result of raw +proof block can be named by following it with +\isakeyword{note}~\isa{some{\isacharunderscore}name\ {\isacharequal}\ this}. As a side effect, +\isa{this} is set to the list of facts on the right-hand side. You +can also say \isa{note\ some{\isacharunderscore}fact}, which simply sets \isa{this}, +i.e.\ recalls \isa{some{\isacharunderscore}fact}, e.g.\ in a \isakeyword{moreover} sequence.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{\isakeyword{fixes}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Sometimes it is necessary to decorate a proposition with type +constraints, as in Cantor's theorem above. These type constraints tend +to make the theorem less readable. The situation can be improved a +little by combining the type constraint with an outer \isa{{\isasymAnd}}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{theorem}\ {\isachardoublequote}{\isasymAnd}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isachardot}\ {\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ f{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent However, now \isa{f} is bound and we need a +\isakeyword{fix}~\isa{f} in the proof before we can refer to \isa{f}. +This is avoided by \isakeyword{fixes}:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{theorem}\ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ f{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent +Even better, \isakeyword{fixes} allows to introduce concrete syntax locally:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ comm{\isacharunderscore}mono{\isacharcolon}\isanewline +\ \ \isakeyword{fixes}\ r\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequote}\ {\isacharparenleft}\isakeyword{infix}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{and}\isanewline +\ \ \ \ \ \ \ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharplus}{\isacharplus}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \isakeyword{assumes}\ comm{\isacharcolon}\ {\isachardoublequote}{\isasymAnd}x\ y{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isachardot}\ x\ {\isacharplus}{\isacharplus}\ y\ {\isacharequal}\ y\ {\isacharplus}{\isacharplus}\ x{\isachardoublequote}\ \isakeyword{and}\isanewline +\ \ \ \ \ \ \ \ \ \ mono{\isacharcolon}\ {\isachardoublequote}{\isasymAnd}x\ y\ z{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isachardot}\ x\ {\isachargreater}\ y\ {\isasymLongrightarrow}\ x\ {\isacharplus}{\isacharplus}\ z\ {\isachargreater}\ y\ {\isacharplus}{\isacharplus}\ z{\isachardoublequote}\isanewline +\ \ \isakeyword{shows}\ {\isachardoublequote}x\ {\isachargreater}\ y\ {\isasymLongrightarrow}\ z\ {\isacharplus}{\isacharplus}\ x\ {\isachargreater}\ z\ {\isacharplus}{\isacharplus}\ y{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ comm\ mono{\isacharparenright}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent The concrete syntax is dropped at the end of the proof and the +theorem becomes \begin{isabelle}% +{\isasymlbrakk}{\isasymAnd}x\ y{\isachardot}\ {\isacharquery}f\ x\ y\ {\isacharequal}\ {\isacharquery}f\ y\ x{\isacharsemicolon}\isanewline +\isaindent{\ \ \ }{\isasymAnd}x\ y\ z{\isachardot}\ {\isacharquery}r\ x\ y\ {\isasymLongrightarrow}\ {\isacharquery}r\ {\isacharparenleft}{\isacharquery}f\ x\ z{\isacharparenright}\ {\isacharparenleft}{\isacharquery}f\ y\ z{\isacharparenright}{\isacharsemicolon}\ {\isacharquery}r\ {\isacharquery}x\ {\isacharquery}y{\isasymrbrakk}\isanewline +{\isasymLongrightarrow}\ {\isacharquery}r\ {\isacharparenleft}{\isacharquery}f\ {\isacharquery}z\ {\isacharquery}x{\isacharparenright}\ {\isacharparenleft}{\isacharquery}f\ {\isacharquery}z\ {\isacharquery}y{\isacharparenright}% +\end{isabelle} +\tweakskip% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{\isakeyword{obtain}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The \isakeyword{obtain} construct can introduce multiple +witnesses and propositions as in the following proof fragment:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ \isakeyword{assumes}\ A{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x\ y{\isachardot}\ P\ x\ y\ {\isasymand}\ Q\ x\ y{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}R{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{from}\ A\ \isamarkupfalse% +\isacommand{obtain}\ x\ y\ \isakeyword{where}\ P{\isacharcolon}\ {\isachardoublequote}P\ x\ y{\isachardoublequote}\ \isakeyword{and}\ Q{\isacharcolon}\ {\isachardoublequote}Q\ x\ y{\isachardoublequote}\ \ \isamarkupfalse% +\isacommand{by}\ blast\isamarkupfalse% +\isamarkupfalse% +% +\begin{isamarkuptext}% +Remember also that one does not even need to start with a formula +containing \isa{{\isasymexists}} as we saw in the proof of Cantor's theorem.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Combining proof styles% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Finally, whole ``scripts'' (tactic-based proofs in the style of +\cite{LNCS2283}) may appear in the leaves of the proof tree, although this is +best avoided. Here is a contrived example:% +\end{isamarkuptext}% +\isamarkuptrue% +\isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ {\isacharparenleft}A\ {\isasymlongrightarrow}\ B{\isacharparenright}\ {\isasymlongrightarrow}\ B{\isachardoublequote}\isanewline +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{assume}\ a{\isacharcolon}\ {\isachardoublequote}A{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}A\ {\isasymlongrightarrow}B{\isacharparenright}\ {\isasymlongrightarrow}\ B{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}rule\ impI{\isacharparenright}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}erule\ impE{\isacharparenright}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}rule\ a{\isacharparenright}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{apply}\ assumption\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{done}\isanewline +\isamarkupfalse% +\isacommand{qed}\isamarkupfalse% +% +\begin{isamarkuptext}% +\noindent You may need to resort to this technique if an +automatic step fails to prove the desired proposition. + +When converting a proof from tactic-style into Isar you can proceed +in a top-down manner: parts of the proof can be left in script form +while the outer structure is already expressed in Isar.% +\end{isamarkuptext}% +\isamarkuptrue% +\isamarkupfalse% +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/intro.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/intro.tex Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,160 @@ +\section{Introduction} + +Isabelle is a generic proof assistant. Isar is an extension of +Isabelle with structured proofs in a stylised language of mathematics. +These proofs are readable for both a human and a machine. +Isabelle/HOL~\cite{LNCS2283} is a specialisation of Isabelle with +higher-order logic (HOL). This paper is a compact introduction to +structured proofs in Isar/HOL, an extension of Isabelle/HOL. We +intentionally do not present the full language but concentrate on the +essentials. Neither do we give a formal semantics of Isar. Instead we +introduce Isar by example. We believe that the language ``speaks for +itself'' in the same way that traditional mathematical proofs do, +which are also introduced by example rather than by teaching students +logic first. A detailed exposition of Isar can be found in Markus +Wenzel's PhD thesis~\cite{Wenzel-PhD} (which also discusses related +work) and the Isar reference manual~\cite{Isar-Ref-Man}. + +\subsection{Background} + +Interactive theorem proving has been dominated by a model of proof +that goes back to the LCF system~\cite{LCF}: a proof is a more or less +structured sequence of commands that manipulate an implicit proof +state. Thus the proof text is only suitable for the machine; for a +human, the proof only comes alive when he can see the state changes +caused by the stepwise execution of the commands. Such proofs are like +uncommented assembly language programs. We call them +\emph{tactic-style} proofs because LCF proof commands were called +tactics. + +A radically different approach was taken by the Mizar +system~\cite{Rudnicki92} where proofs are written in a stylised language akin +to that used in ordinary mathematics texts. The most important argument in +favour of a mathematics-like proof language is communication: as soon as not +just the theorem but also the proof becomes an object of interest, it should +be readable. From a system development point of view there is a second +important argument against tactic-style proofs: they are much harder to +maintain when the system is modified. +%The reason is as follows. Since it is +%usually quite unclear what exactly is proved at some point in the middle of a +%command sequence, updating a failed proof often requires executing the proof +%up to the point of failure using the old version of the system. In contrast, +%mathematics-like proofs contain enough information to tell you what is proved +%at any point. + +For these reasons the Isabelle system, originally firmly in the +LCF-tradition, was extended with a language for writing structured +proofs in a mathematics-like style. As the name already indicates, +Isar was certainly inspired by Mizar. However, there are many +differences. For a start, Isar is generic: only a few of the language +constructs described below are specific to HOL; many are generic and +thus available for any logic defined in Isabelle, e.g.\ ZF. +Furthermore, we have Isabelle's powerful automatic proof procedures at +our disposal. A closer comparison of Isar and Mizar can be found +elsewhere~\cite{WenzelW-JAR}. + +\subsection{A first glimpse of Isar} +Below you find a simplified grammar for Isar proofs. +Parentheses are used for grouping and $^?$ indicates an optional item: +\begin{center} +\begin{tabular}{lrl} +\emph{proof} & ::= & \isakeyword{proof} \emph{method}$^?$ + \emph{statement}* + \isakeyword{qed} \\ + &$\mid$& \isakeyword{by} \emph{method}\\[1ex] +\emph{statement} &::= & \isakeyword{fix} \emph{variables} \\ + &$\mid$& \isakeyword{assume} \emph{propositions} \\ + &$\mid$& (\isakeyword{from} \emph{facts})$^?$ + (\isakeyword{show} $\mid$ \isakeyword{have}) + \emph{propositions} \emph{proof} \\[1ex] +\emph{proposition} &::=& (\emph{label}{\bf:})$^?$ \emph{string} \\[1ex] +\emph{fact} &::=& \emph{label} +\end{tabular} +\end{center} +A proof can be either compound (\isakeyword{proof} -- +\isakeyword{qed}) or atomic (\isakeyword{by}). A \emph{method} is a +proof method (tactic) offered by the underlying theorem prover. +Thus this grammar is generic both w.r.t.\ the logic and the theorem prover. + +This is a typical proof skeleton: +\begin{center} +\begin{tabular}{@{}ll} +\isakeyword{proof}\\ +\hspace*{3ex}\isakeyword{assume} \isa{"}\emph{the-assm}\isa{"}\\ +\hspace*{3ex}\isakeyword{have} \isa{"}\dots\isa{"} & --- intermediate result\\ +\hspace*{3ex}\vdots\\ +\hspace*{3ex}\isakeyword{have} \isa{"}\dots\isa{"} & --- intermediate result\\ +\hspace*{3ex}\isakeyword{show} \isa{"}\emph{the-concl}\isa{"}\\ +\isakeyword{qed} +\end{tabular} +\end{center} +It proves \emph{the-assm}~$\Longrightarrow$~\emph{the-concl}. Text starting with +``---'' is a comment. The intermediate \isakeyword{have}s are only +there to bridge the gap between the assumption and the conclusion and +do not contribute to the theorem being proved. In contrast, +\isakeyword{show} establishes the conclusion of the theorem. + +\subsection{Bits of Isabelle} + +We recall some basic notions and notation from Isabelle. For more +details and for instructions how to run examples see +elsewhere~\cite{LNCS2283}. + +Isabelle's meta-logic comes with a type of \emph{propositions} with +implication $\Longrightarrow$ and a universal quantifier $\bigwedge$ for expressing +inference rules and generality. Iterated implications $A_1 \Longrightarrow \dots +A_n \Longrightarrow A$ may be abbreviated to $[\![ A_1; \dots; A_n ]\!] \Longrightarrow A$. +Applying a theorem $A \Longrightarrow B$ (named \isa{T}) to a theorem $A$ (named +\isa{U}) is written \isa{T[OF U]} and yields theorem $B$. + +Isabelle terms are simply typed. Function types are +written $\tau_1 \Rightarrow \tau_2$. + +Free variables that may be instantiated (``logical variables'' in Prolog +parlance) are prefixed with a \isa{?}. Typically, theorems are stated with +ordinary free variables but after the proof those are automatically replaced +by \isa{?}-variables. Thus the theorem can be used with arbitrary instances +of its free variables. + +Isabelle/HOL offers all the usual logical symbols like $\longrightarrow$, $\land$, +$\forall$ etc. HOL formulae are propositions, e.g.\ $\forall$ can appear below +$\Longrightarrow$, but not the other way around. Beware that $\longrightarrow$ binds more +tightly than $\Longrightarrow$: in $\forall x. P \longrightarrow Q$ the $\forall x$ covers $P \longrightarrow Q$, whereas +in $\forall x. P \Longrightarrow Q$ it covers only $P$. + +Proof methods include \isa{rule} (which performs a backwards +step with a given rule, unifying the conclusion of the rule with the +current subgoal and replacing the subgoal by the premises of the +rule), \isa{simp} (for simplification) and \isa{blast} (for predicate +calculus reasoning). + +\subsection{Overview of the paper} + +The rest of the paper is divided into two parts. +Section~\ref{sec:Logic} introduces proofs in pure logic based on +natural deduction. Section~\ref{sec:Induct} is dedicated to induction, +the key reasoning principle for computer science applications. + +There are two further areas where Isar provides specific support, but +which we do not document here. Reasoning by chains of (in)equations is +described elsewhere~\cite{BauerW-TPHOLs01}. Reasoning about +axiomatically defined structures by means of so called ``locales'' was +first described in \cite{KWP-TPHOLs99} but has evolved much since +then. + +Finally, a word of warning for potential writers of Isar proofs. It +has always been easier to write obscure rather than readable texts. +Similarly, tactic-style proofs are often (though by no means always!) +easier to write than readable ones: structure does not emerge +automatically but needs to be understood and imposed. If the precise +structure of the proof is unclear at beginning, it can be useful to +start in a tactic-based style for exploratory purposes until one has +found a proof which can be converted into a structured text in a +second step. +% Top down conversion is possible because Isar +%allows tactic-style proofs as components of structured ones. + +%Do not be mislead by the simplicity of the formulae being proved, +%especially in the beginning. Isar has been used very successfully in +%large applications, for example the formalisation of Java, in +%particular the verification of the bytecode verifier~\cite{KleinN-TCS}. diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/isabelle.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/isabelle.sty Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,157 @@ +%% +%% Author: Markus Wenzel, TU Muenchen +%% License: GPL (GNU GENERAL PUBLIC LICENSE) +%% +%% macros for Isabelle generated LaTeX output +%% + +%%% Simple document preparation (based on theory token language and symbols) + +% isabelle environments + +\newcommand{\isabellecontext}{UNKNOWN} + +\newcommand{\isastyle}{\small\tt\slshape} +\newcommand{\isastyleminor}{\small\tt\slshape} +\newcommand{\isastylescript}{\footnotesize\tt\slshape} +\newcommand{\isastyletext}{\normalsize\rm} +\newcommand{\isastyletxt}{\rm} +\newcommand{\isastylecmt}{\rm} + +%symbol markup -- \emph achieves decent spacing via italic corrections +\newcommand{\isamath}[1]{\emph{$#1$}} +\newcommand{\isatext}[1]{\emph{#1}} +\newcommand{\isascriptstyle}{\def\isamath##1{##1}\def\isatext##1{\mbox{\isastylescript##1}}} +\newcommand{\isactrlsub}[1]{\emph{\isascriptstyle${}\sb{#1}$}} +\newcommand{\isactrlsup}[1]{\emph{\isascriptstyle${}\sp{#1}$}} +\newcommand{\isactrlbold}[1]{{\bfseries\upshape\boldmath#1}} + +\newdimen\isa@parindent\newdimen\isa@parskip + +\newenvironment{isabellebody}{% +\isamarkuptrue\par% +\isa@parindent\parindent\parindent0pt% +\isa@parskip\parskip\parskip0pt% +\isastyle}{\par} + +\newenvironment{isabelle} +{\begin{trivlist}\begin{isabellebody}\item\relax} +{\end{isabellebody}\end{trivlist}} + +\newcommand{\isa}[1]{\emph{\isastyleminor #1}} + +\newcommand{\isaindent}[1]{\hphantom{#1}} +\newcommand{\isanewline}{\mbox{}\par\mbox{}} +\newcommand{\isadigit}[1]{#1} + +\chardef\isacharbang=`\! +\chardef\isachardoublequote=`\" +\chardef\isacharhash=`\# +\chardef\isachardollar=`\$ +\chardef\isacharpercent=`\% +\chardef\isacharampersand=`\& +\chardef\isacharprime=`\' +\chardef\isacharparenleft=`\( +\chardef\isacharparenright=`\) +\chardef\isacharasterisk=`\* +\chardef\isacharplus=`\+ +\chardef\isacharcomma=`\, +\chardef\isacharminus=`\- +\chardef\isachardot=`\. +\chardef\isacharslash=`\/ +\chardef\isacharcolon=`\: +\chardef\isacharsemicolon=`\; +\chardef\isacharless=`\< +\chardef\isacharequal=`\= +\chardef\isachargreater=`\> +\chardef\isacharquery=`\? +\chardef\isacharat=`\@ +\chardef\isacharbrackleft=`\[ +\chardef\isacharbackslash=`\\ +\chardef\isacharbrackright=`\] +\chardef\isacharcircum=`\^ +\chardef\isacharunderscore=`\_ +\chardef\isacharbackquote=`\` +\chardef\isacharbraceleft=`\{ +\chardef\isacharbar=`\| +\chardef\isacharbraceright=`\} +\chardef\isachartilde=`\~ + + +% keyword and section markup + +\newcommand{\isakeywordcharunderscore}{\_} +\newcommand{\isakeyword}[1] +{\emph{\bf\def\isachardot{.}\def\isacharunderscore{\isakeywordcharunderscore}% +\def\isacharbraceleft{\{}\def\isacharbraceright{\}}#1}} +\newcommand{\isacommand}[1]{\isakeyword{#1}} + +\newcommand{\isamarkupheader}[1]{\section{#1}} +\newcommand{\isamarkupchapter}[1]{\chapter{#1}} +\newcommand{\isamarkupsection}[1]{\section{#1}} +\newcommand{\isamarkupsubsection}[1]{\subsection{#1}} +\newcommand{\isamarkupsubsubsection}[1]{\subsubsection{#1}} +\newcommand{\isamarkupsect}[1]{\section{#1}} +\newcommand{\isamarkupsubsect}[1]{\subsection{#1}} +\newcommand{\isamarkupsubsubsect}[1]{\subsubsection{#1}} + +\newif\ifisamarkup +\newcommand{\isabeginpar}{\par\ifisamarkup\relax\else\medskip\fi} +\newcommand{\isaendpar}{\par\medskip} +\newenvironment{isapar}{\parindent\isa@parindent\parskip\isa@parskip\isabeginpar}{\isaendpar} +\newenvironment{isamarkuptext}{\isastyletext\begin{isapar}}{\end{isapar}} +\newenvironment{isamarkuptxt}{\isastyletxt\begin{isapar}}{\end{isapar}} +\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}} + + +% alternative styles + +\newcommand{\isabellestyle}{} +\def\isabellestyle#1{\csname isabellestyle#1\endcsname} + +\newcommand{\isabellestylett}{% +\renewcommand{\isastyle}{\small\tt}% +\renewcommand{\isastyleminor}{\small\tt}% +\renewcommand{\isastylescript}{\footnotesize\tt}% +} +\newcommand{\isabellestyleit}{% +\renewcommand{\isastyle}{\small\it}% +\renewcommand{\isastyleminor}{\it}% +\renewcommand{\isastylescript}{\footnotesize\it}% +\renewcommand{\isakeywordcharunderscore}{\mbox{-}}% +\renewcommand{\isacharbang}{\isamath{!}}% +\renewcommand{\isachardoublequote}{}% +\renewcommand{\isacharhash}{\isamath{\#}}% +\renewcommand{\isachardollar}{\isamath{\$}}% +\renewcommand{\isacharpercent}{\isamath{\%}}% +\renewcommand{\isacharampersand}{\isamath{\&}}% +\renewcommand{\isacharprime}{\isamath{\mskip2mu{'}\mskip-2mu}}% +\renewcommand{\isacharparenleft}{\isamath{(}}% +\renewcommand{\isacharparenright}{\isamath{)}}% +\renewcommand{\isacharasterisk}{\isamath{*}}% +\renewcommand{\isacharplus}{\isamath{+}}% +\renewcommand{\isacharcomma}{\isamath{\mathord,}}% +\renewcommand{\isacharminus}{\isamath{-}}% +\renewcommand{\isachardot}{\isamath{\mathord.}}% +\renewcommand{\isacharslash}{\isamath{/}}% +\renewcommand{\isacharcolon}{\isamath{\mathord:}}% +\renewcommand{\isacharsemicolon}{\isamath{\mathord;}}% +\renewcommand{\isacharless}{\isamath{<}}% +\renewcommand{\isacharequal}{\isamath{=}}% +\renewcommand{\isachargreater}{\isamath{>}}% +\renewcommand{\isacharat}{\isamath{@}}% +\renewcommand{\isacharbrackleft}{\isamath{[}}% +\renewcommand{\isacharbrackright}{\isamath{]}}% +\renewcommand{\isacharunderscore}{\mbox{-}}% +\renewcommand{\isacharbraceleft}{\isamath{\{}}% +\renewcommand{\isacharbar}{\isamath{\mid}}% +\renewcommand{\isacharbraceright}{\isamath{\}}}% +\renewcommand{\isachartilde}{\isamath{{}\sp{\sim}}}% +} + +\newcommand{\isabellestylesl}{% +\isabellestyleit% +\renewcommand{\isastyle}{\small\sl}% +\renewcommand{\isastyleminor}{\sl}% +\renewcommand{\isastylescript}{\footnotesize\sl}% +} diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/isabellesym.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/isabellesym.sty Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,354 @@ +%% +%% Author: Markus Wenzel, TU Muenchen +%% License: GPL (GNU GENERAL PUBLIC LICENSE) +%% +%% definitions of standard Isabelle symbols +%% + +% symbol definitions + +\newcommand{\isasymzero}{\isamath{\mathbf{0}}} %requires amssym +\newcommand{\isasymone}{\isamath{\mathbf{1}}} %requires amssym +\newcommand{\isasymtwo}{\isamath{\mathbf{2}}} %requires amssym +\newcommand{\isasymthree}{\isamath{\mathbf{3}}} %requires amssym +\newcommand{\isasymfour}{\isamath{\mathbf{4}}} %requires amssym +\newcommand{\isasymfive}{\isamath{\mathbf{5}}} %requires amssym +\newcommand{\isasymsix}{\isamath{\mathbf{6}}} %requires amssym +\newcommand{\isasymseven}{\isamath{\mathbf{7}}} %requires amssym +\newcommand{\isasymeight}{\isamath{\mathbf{8}}} %requires amssym +\newcommand{\isasymnine}{\isamath{\mathbf{9}}} %requires amssym +\newcommand{\isasymA}{\isamath{\mathcal{A}}} +\newcommand{\isasymB}{\isamath{\mathcal{B}}} +\newcommand{\isasymC}{\isamath{\mathcal{C}}} +\newcommand{\isasymD}{\isamath{\mathcal{D}}} +\newcommand{\isasymE}{\isamath{\mathcal{E}}} +\newcommand{\isasymF}{\isamath{\mathcal{F}}} +\newcommand{\isasymG}{\isamath{\mathcal{G}}} +\newcommand{\isasymH}{\isamath{\mathcal{H}}} +\newcommand{\isasymI}{\isamath{\mathcal{I}}} +\newcommand{\isasymJ}{\isamath{\mathcal{J}}} +\newcommand{\isasymK}{\isamath{\mathcal{K}}} +\newcommand{\isasymL}{\isamath{\mathcal{L}}} +\newcommand{\isasymM}{\isamath{\mathcal{M}}} +\newcommand{\isasymN}{\isamath{\mathcal{N}}} +\newcommand{\isasymO}{\isamath{\mathcal{O}}} +\newcommand{\isasymP}{\isamath{\mathcal{P}}} +\newcommand{\isasymQ}{\isamath{\mathcal{Q}}} +\newcommand{\isasymR}{\isamath{\mathcal{R}}} +\newcommand{\isasymS}{\isamath{\mathcal{S}}} +\newcommand{\isasymT}{\isamath{\mathcal{T}}} +\newcommand{\isasymU}{\isamath{\mathcal{U}}} +\newcommand{\isasymV}{\isamath{\mathcal{V}}} +\newcommand{\isasymW}{\isamath{\mathcal{W}}} +\newcommand{\isasymX}{\isamath{\mathcal{X}}} +\newcommand{\isasymY}{\isamath{\mathcal{Y}}} +\newcommand{\isasymZ}{\isamath{\mathcal{Z}}} +\newcommand{\isasyma}{\isamath{\mathrm{a}}} +\newcommand{\isasymb}{\isamath{\mathrm{b}}} +\newcommand{\isasymc}{\isamath{\mathrm{c}}} +\newcommand{\isasymd}{\isamath{\mathrm{d}}} +\newcommand{\isasyme}{\isamath{\mathrm{e}}} +\newcommand{\isasymf}{\isamath{\mathrm{f}}} +\newcommand{\isasymg}{\isamath{\mathrm{g}}} +\newcommand{\isasymh}{\isamath{\mathrm{h}}} +\newcommand{\isasymi}{\isamath{\mathrm{i}}} +\newcommand{\isasymj}{\isamath{\mathrm{j}}} +\newcommand{\isasymk}{\isamath{\mathrm{k}}} +\newcommand{\isasyml}{\isamath{\mathrm{l}}} +\newcommand{\isasymm}{\isamath{\mathrm{m}}} +\newcommand{\isasymn}{\isamath{\mathrm{n}}} +\newcommand{\isasymo}{\isamath{\mathrm{o}}} +\newcommand{\isasymp}{\isamath{\mathrm{p}}} +\newcommand{\isasymq}{\isamath{\mathrm{q}}} +\newcommand{\isasymr}{\isamath{\mathrm{r}}} +\newcommand{\isasyms}{\isamath{\mathrm{s}}} +\newcommand{\isasymt}{\isamath{\mathrm{t}}} +\newcommand{\isasymu}{\isamath{\mathrm{u}}} +\newcommand{\isasymv}{\isamath{\mathrm{v}}} +\newcommand{\isasymw}{\isamath{\mathrm{w}}} +\newcommand{\isasymx}{\isamath{\mathrm{x}}} +\newcommand{\isasymy}{\isamath{\mathrm{y}}} +\newcommand{\isasymz}{\isamath{\mathrm{z}}} +\newcommand{\isasymAA}{\isamath{\mathfrak{A}}} %requires eufrak +\newcommand{\isasymBB}{\isamath{\mathfrak{B}}} %requires eufrak +\newcommand{\isasymCC}{\isamath{\mathfrak{C}}} %requires eufrak +\newcommand{\isasymDD}{\isamath{\mathfrak{D}}} %requires eufrak +\newcommand{\isasymEE}{\isamath{\mathfrak{E}}} %requires eufrak +\newcommand{\isasymFF}{\isamath{\mathfrak{F}}} %requires eufrak +\newcommand{\isasymGG}{\isamath{\mathfrak{G}}} %requires eufrak +\newcommand{\isasymHH}{\isamath{\mathfrak{H}}} %requires eufrak +\newcommand{\isasymII}{\isamath{\mathfrak{I}}} %requires eufrak +\newcommand{\isasymJJ}{\isamath{\mathfrak{J}}} %requires eufrak +\newcommand{\isasymKK}{\isamath{\mathfrak{K}}} %requires eufrak +\newcommand{\isasymLL}{\isamath{\mathfrak{L}}} %requires eufrak +\newcommand{\isasymMM}{\isamath{\mathfrak{M}}} %requires eufrak +\newcommand{\isasymNN}{\isamath{\mathfrak{N}}} %requires eufrak +\newcommand{\isasymOO}{\isamath{\mathfrak{O}}} %requires eufrak +\newcommand{\isasymPP}{\isamath{\mathfrak{P}}} %requires eufrak +\newcommand{\isasymQQ}{\isamath{\mathfrak{Q}}} %requires eufrak +\newcommand{\isasymRR}{\isamath{\mathfrak{R}}} %requires eufrak +\newcommand{\isasymSS}{\isamath{\mathfrak{S}}} %requires eufrak +\newcommand{\isasymTT}{\isamath{\mathfrak{T}}} %requires eufrak +\newcommand{\isasymUU}{\isamath{\mathfrak{U}}} %requires eufrak +\newcommand{\isasymVV}{\isamath{\mathfrak{V}}} %requires eufrak +\newcommand{\isasymWW}{\isamath{\mathfrak{W}}} %requires eufrak +\newcommand{\isasymXX}{\isamath{\mathfrak{X}}} %requires eufrak +\newcommand{\isasymYY}{\isamath{\mathfrak{Y}}} %requires eufrak +\newcommand{\isasymZZ}{\isamath{\mathfrak{Z}}} %requires eufrak +\newcommand{\isasymaa}{\isamath{\mathfrak{a}}} %requires eufrak +\newcommand{\isasymbb}{\isamath{\mathfrak{b}}} %requires eufrak +\newcommand{\isasymcc}{\isamath{\mathfrak{c}}} %requires eufrak +\newcommand{\isasymdd}{\isamath{\mathfrak{d}}} %requires eufrak +\newcommand{\isasymee}{\isamath{\mathfrak{e}}} %requires eufrak +\newcommand{\isasymff}{\isamath{\mathfrak{f}}} %requires eufrak +\newcommand{\isasymgg}{\isamath{\mathfrak{g}}} %requires eufrak +\newcommand{\isasymhh}{\isamath{\mathfrak{h}}} %requires eufrak +\newcommand{\isasymii}{\isamath{\mathfrak{i}}} %requires eufrak +\newcommand{\isasymjj}{\isamath{\mathfrak{j}}} %requires eufrak +\newcommand{\isasymkk}{\isamath{\mathfrak{k}}} %requires eufrak +\newcommand{\isasymll}{\isamath{\mathfrak{l}}} %requires eufrak +\newcommand{\isasymmm}{\isamath{\mathfrak{m}}} %requires eufrak +\newcommand{\isasymnn}{\isamath{\mathfrak{n}}} %requires eufrak +\newcommand{\isasymoo}{\isamath{\mathfrak{o}}} %requires eufrak +\newcommand{\isasympp}{\isamath{\mathfrak{p}}} %requires eufrak +\newcommand{\isasymqq}{\isamath{\mathfrak{q}}} %requires eufrak +\newcommand{\isasymrr}{\isamath{\mathfrak{r}}} %requires eufrak +\newcommand{\isasymss}{\isamath{\mathfrak{s}}} %requires eufrak +\newcommand{\isasymtt}{\isamath{\mathfrak{t}}} %requires eufrak +\newcommand{\isasymuu}{\isamath{\mathfrak{u}}} %requires eufrak +\newcommand{\isasymvv}{\isamath{\mathfrak{v}}} %requires eufrak +\newcommand{\isasymww}{\isamath{\mathfrak{w}}} %requires eufrak +\newcommand{\isasymxx}{\isamath{\mathfrak{x}}} %requires eufrak +\newcommand{\isasymyy}{\isamath{\mathfrak{y}}} %requires eufrak +\newcommand{\isasymzz}{\isamath{\mathfrak{z}}} %requires eufrak +\newcommand{\isasymalpha}{\isamath{\alpha}} +\newcommand{\isasymbeta}{\isamath{\beta}} +\newcommand{\isasymgamma}{\isamath{\gamma}} +\newcommand{\isasymdelta}{\isamath{\delta}} +\newcommand{\isasymepsilon}{\isamath{\varepsilon}} +\newcommand{\isasymzeta}{\isamath{\zeta}} +\newcommand{\isasymeta}{\isamath{\eta}} +\newcommand{\isasymtheta}{\isamath{\vartheta}} +\newcommand{\isasymiota}{\isamath{\iota}} +\newcommand{\isasymkappa}{\isamath{\kappa}} +\newcommand{\isasymlambda}{\isamath{\lambda}} +\newcommand{\isasymmu}{\isamath{\mu}} +\newcommand{\isasymnu}{\isamath{\nu}} +\newcommand{\isasymxi}{\isamath{\xi}} +\newcommand{\isasympi}{\isamath{\pi}} +\newcommand{\isasymrho}{\isamath{\varrho}} +\newcommand{\isasymsigma}{\isamath{\sigma}} +\newcommand{\isasymtau}{\isamath{\tau}} +\newcommand{\isasymupsilon}{\isamath{\upsilon}} +\newcommand{\isasymphi}{\isamath{\varphi}} +\newcommand{\isasymchi}{\isamath{\chi}} +\newcommand{\isasympsi}{\isamath{\psi}} +\newcommand{\isasymomega}{\isamath{\omega}} +\newcommand{\isasymGamma}{\isamath{\Gamma}} +\newcommand{\isasymDelta}{\isamath{\Delta}} +\newcommand{\isasymTheta}{\isamath{\Theta}} +\newcommand{\isasymLambda}{\isamath{\Lambda}} +\newcommand{\isasymXi}{\isamath{\Xi}} +\newcommand{\isasymPi}{\isamath{\Pi}} +\newcommand{\isasymSigma}{\isamath{\Sigma}} +\newcommand{\isasymUpsilon}{\isamath{\Upsilon}} +\newcommand{\isasymPhi}{\isamath{\Phi}} +\newcommand{\isasymPsi}{\isamath{\Psi}} +\newcommand{\isasymOmega}{\isamath{\Omega}} +\newcommand{\isasymbool}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{B}}} +\newcommand{\isasymcomplex}{\isamath{\mathrm{C}\mkern-15mu{\phantom{\mathrm{t}}\vrule}\mkern9mu}} +\newcommand{\isasymnat}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{N}}} +\newcommand{\isasymrat}{\isamath{\mathrm{Q}\mkern-16mu{\phantom{\mathrm{t}}\vrule}\mkern10mu}} +\newcommand{\isasymreal}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{R}}} +\newcommand{\isasymint}{\isamath{\mathsf{Z}\mkern-7.5mu\mathsf{Z}}} +\newcommand{\isasymleftarrow}{\isamath{\leftarrow}} +\newcommand{\isasymlongleftarrow}{\isamath{\longleftarrow}} +\newcommand{\isasymrightarrow}{\isamath{\rightarrow}} +\newcommand{\isasymlongrightarrow}{\isamath{\longrightarrow}} +\newcommand{\isasymLeftarrow}{\isamath{\Leftarrow}} +\newcommand{\isasymLongleftarrow}{\isamath{\Longleftarrow}} +\newcommand{\isasymRightarrow}{\isamath{\Rightarrow}} +\newcommand{\isasymLongrightarrow}{\isamath{\Longrightarrow}} +\newcommand{\isasymleftrightarrow}{\isamath{\leftrightarrow}} +\newcommand{\isasymlongleftrightarrow}{\isamath{\longleftrightarrow}} +\newcommand{\isasymLeftrightarrow}{\isamath{\Leftrightarrow}} +\newcommand{\isasymLongleftrightarrow}{\isamath{\Longleftrightarrow}} +\newcommand{\isasymmapsto}{\isamath{\mapsto}} +\newcommand{\isasymlongmapsto}{\isamath{\longmapsto}} +\newcommand{\isasymmidarrow}{\isamath{\relbar}} +\newcommand{\isasymMidarrow}{\isamath{\Relbar}} +\newcommand{\isasymhookleftarrow}{\isamath{\hookleftarrow}} +\newcommand{\isasymhookrightarrow}{\isamath{\hookrightarrow}} +\newcommand{\isasymleftharpoondown}{\isamath{\leftharpoondown}} +\newcommand{\isasymrightharpoondown}{\isamath{\rightharpoondown}} +\newcommand{\isasymleftharpoonup}{\isamath{\leftharpoonup}} +\newcommand{\isasymrightharpoonup}{\isamath{\rightharpoonup}} +\newcommand{\isasymrightleftharpoons}{\isamath{\rightleftharpoons}} +\newcommand{\isasymleadsto}{\isamath{\leadsto}} %requires amssym +\newcommand{\isasymup}{\isamath{\uparrow}} +\newcommand{\isasymUp}{\isamath{\Uparrow}} +\newcommand{\isasymdown}{\isamath{\downarrow}} +\newcommand{\isasymDown}{\isamath{\Downarrow}} +\newcommand{\isasymupdown}{\isamath{\updownarrow}} +\newcommand{\isasymUpdown}{\isamath{\Updownarrow}} +\newcommand{\isasymlangle}{\isamath{\langle}} +\newcommand{\isasymrangle}{\isamath{\rangle}} +\newcommand{\isasymlceil}{\isamath{\lceil}} +\newcommand{\isasymrceil}{\isamath{\rceil}} +\newcommand{\isasymlfloor}{\isamath{\lfloor}} +\newcommand{\isasymrfloor}{\isamath{\rfloor}} +\newcommand{\isasymlparr}{\isamath{\mathopen{(\mkern-3mu\mid}}} +\newcommand{\isasymrparr}{\isamath{\mathclose{\mid\mkern-3mu)}}} +\newcommand{\isasymlbrakk}{\isamath{\mathopen{\lbrack\mkern-3mu\lbrack}}} +\newcommand{\isasymrbrakk}{\isamath{\mathclose{\rbrack\mkern-3mu\rbrack}}} +\newcommand{\isasymlbrace}{\isamath{\mathopen{\lbrace\mkern-4.5mu\mid}}} +\newcommand{\isasymrbrace}{\isamath{\mathclose{\mid\mkern-4.5mu\rbrace}}} +\newcommand{\isasymguillemotleft}{\isatext{\flqq}} %requires babel +\newcommand{\isasymguillemotright}{\isatext{\frqq}} %requires babel +\newcommand{\isasymColon}{\isamath{\mathrel{::}}} +\newcommand{\isasymnot}{\isamath{\neg}} +\newcommand{\isasymbottom}{\isamath{\bot}} +\newcommand{\isasymtop}{\isamath{\top}} +\newcommand{\isasymand}{\isamath{\wedge}} +\newcommand{\isasymAnd}{\isamath{\bigwedge}} +\newcommand{\isasymor}{\isamath{\vee}} +\newcommand{\isasymOr}{\isamath{\bigvee}} +\newcommand{\isasymforall}{\isamath{\forall\,}} +\newcommand{\isasymexists}{\isamath{\exists\,}} +\newcommand{\isasymbox}{\isamath{\Box}} %requires amssym +\newcommand{\isasymdiamond}{\isamath{\Diamond}} %requires amssym +\newcommand{\isasymturnstile}{\isamath{\vdash}} +\newcommand{\isasymTurnstile}{\isamath{\models}} +\newcommand{\isasymtturnstile}{\isamath{\vdash\!\!\!\vdash}} +\newcommand{\isasymTTurnstile}{\isamath{\mid\!\models}} +\newcommand{\isasymstileturn}{\isamath{\dashv}} +\newcommand{\isasymsurd}{\isamath{\surd}} +\newcommand{\isasymle}{\isamath{\le}} +\newcommand{\isasymge}{\isamath{\ge}} +\newcommand{\isasymlless}{\isamath{\ll}} +\newcommand{\isasymggreater}{\isamath{\gg}} +\newcommand{\isasymlesssim}{\isamath{\lesssim}} %requires amssymb +\newcommand{\isasymgreatersim}{\isamath{\gtrsim}} %requires amssymb +\newcommand{\isasymlessapprox}{\isamath{\lessapprox}} %requires amssymb +\newcommand{\isasymgreaterapprox}{\isamath{\gtrapprox}} %requires amssymb +\newcommand{\isasymin}{\isamath{\in}} +\newcommand{\isasymnotin}{\isamath{\notin}} +\newcommand{\isasymsubset}{\isamath{\subset}} +\newcommand{\isasymsupset}{\isamath{\supset}} +\newcommand{\isasymsubseteq}{\isamath{\subseteq}} +\newcommand{\isasymsupseteq}{\isamath{\supseteq}} +\newcommand{\isasymsqsubset}{\isamath{\sqsubset}} +\newcommand{\isasymsqsupset}{\isamath{\sqsupset}} %requires amssym +\newcommand{\isasymsqsubseteq}{\isamath{\sqsubseteq}} +\newcommand{\isasymsqsupseteq}{\isamath{\sqsupseteq}} +\newcommand{\isasyminter}{\isamath{\cap}} +\newcommand{\isasymInter}{\isamath{\bigcap\,}} +\newcommand{\isasymunion}{\isamath{\cup}} +\newcommand{\isasymUnion}{\isamath{\bigcup\,}} +\newcommand{\isasymsqunion}{\isamath{\sqcup}} +\newcommand{\isasymSqunion}{\isamath{\bigsqcup\,}} +\newcommand{\isasymsqinter}{\isamath{\sqcap}} +\newcommand{\isasymSqinter}{\isamath{\bigsqcap\,}} %requires masmath +\newcommand{\isasymuplus}{\isamath{\uplus}} +\newcommand{\isasymUplus}{\isamath{\biguplus\,}} +\newcommand{\isasymnoteq}{\isamath{\not=}} +\newcommand{\isasymsim}{\isamath{\sim}} +\newcommand{\isasymdoteq}{\isamath{\doteq}} +\newcommand{\isasymsimeq}{\isamath{\simeq}} +\newcommand{\isasymapprox}{\isamath{\approx}} +\newcommand{\isasymasymp}{\isamath{\asymp}} +\newcommand{\isasymcong}{\isamath{\cong}} +\newcommand{\isasymsmile}{\isamath{\smile}} +\newcommand{\isasymequiv}{\isamath{\equiv}} +\newcommand{\isasymfrown}{\isamath{\frown}} +\newcommand{\isasympropto}{\isamath{\propto}} +\newcommand{\isasymbowtie}{\isamath{\bowtie}} +\newcommand{\isasymprec}{\isamath{\prec}} +\newcommand{\isasymsucc}{\isamath{\succ}} +\newcommand{\isasympreceq}{\isamath{\preceq}} +\newcommand{\isasymsucceq}{\isamath{\succeq}} +\newcommand{\isasymparallel}{\isamath{\parallel}} +\newcommand{\isasymbar}{\isamath{\mid}} +\newcommand{\isasymplusminus}{\isamath{\pm}} +\newcommand{\isasymminusplus}{\isamath{\mp}} +\newcommand{\isasymtimes}{\isamath{\times}} +\newcommand{\isasymdiv}{\isamath{\div}} +\newcommand{\isasymcdot}{\isamath{\cdot}} +\newcommand{\isasymstar}{\isamath{\star}} +\newcommand{\isasymbullet}{\boldmath\isamath{\mathchoice{\displaystyle{\cdot}}{\textstyle{\cdot}}{\scriptstyle{\bullet}}{\scriptscriptstyle{\bullet}}}} +\newcommand{\isasymcirc}{\isamath{\circ}} +\newcommand{\isasymdagger}{\isamath{\dagger}} +\newcommand{\isasymddagger}{\isamath{\ddagger}} +\newcommand{\isasymlhd}{\isamath{\lhd}} +\newcommand{\isasymrhd}{\isamath{\rhd}} +\newcommand{\isasymunlhd}{\isamath{\unlhd}} +\newcommand{\isasymunrhd}{\isamath{\unrhd}} +\newcommand{\isasymtriangleleft}{\isamath{\triangleleft}} +\newcommand{\isasymtriangleright}{\isamath{\triangleright}} +\newcommand{\isasymtriangle}{\isamath{\triangle}} +\newcommand{\isasymtriangleq}{\isamath{\triangleq}} %requires amssymb +\newcommand{\isasymoplus}{\isamath{\oplus}} +\newcommand{\isasymOplus}{\isamath{\bigoplus\,}} +\newcommand{\isasymotimes}{\isamath{\otimes}} +\newcommand{\isasymOtimes}{\isamath{\bigotimes\,}} +\newcommand{\isasymodot}{\isamath{\odot}} +\newcommand{\isasymOdot}{\isamath{\bigodot\,}} +\newcommand{\isasymominus}{\isamath{\ominus}} +\newcommand{\isasymoslash}{\isamath{\oslash}} +\newcommand{\isasymdots}{\isamath{\dots}} +\newcommand{\isasymcdots}{\isamath{\cdots}} +\newcommand{\isasymSum}{\isamath{\sum\,}} +\newcommand{\isasymProd}{\isamath{\prod\,}} +\newcommand{\isasymCoprod}{\isamath{\coprod\,}} +\newcommand{\isasyminfinity}{\isamath{\infty}} +\newcommand{\isasymintegral}{\isamath{\int\,}} +\newcommand{\isasymointegral}{\isamath{\oint\,}} +\newcommand{\isasymclubsuit}{\isamath{\clubsuit}} +\newcommand{\isasymdiamondsuit}{\isamath{\diamondsuit}} +\newcommand{\isasymheartsuit}{\isamath{\heartsuit}} +\newcommand{\isasymspadesuit}{\isamath{\spadesuit}} +\newcommand{\isasymaleph}{\isamath{\aleph}} +\newcommand{\isasymemptyset}{\isamath{\emptyset}} +\newcommand{\isasymnabla}{\isamath{\nabla}} +\newcommand{\isasympartial}{\isamath{\partial}} +\newcommand{\isasymRe}{\isamath{\Re}} +\newcommand{\isasymIm}{\isamath{\Im}} +\newcommand{\isasymflat}{\isamath{\flat}} +\newcommand{\isasymnatural}{\isamath{\natural}} +\newcommand{\isasymsharp}{\isamath{\sharp}} +\newcommand{\isasymangle}{\isamath{\angle}} +\newcommand{\isasymcopyright}{\isatext{\rm\copyright}} +\newcommand{\isasymregistered}{\isatext{\rm\textregistered}} +\newcommand{\isasymhyphen}{\isatext{\rm-}} +\newcommand{\isasyminverse}{\isamath{{}^{-1}}} +\newcommand{\isasymonesuperior}{\isamath{\mathonesuperior}} %requires latin1 +\newcommand{\isasymonequarter}{\isatext{\rm\textonequarter}} %requires latin1 +\newcommand{\isasymtwosuperior}{\isamath{\mathtwosuperior}} %requires latin1 +\newcommand{\isasymonehalf}{\isatext{\rm\textonehalf}} %requires latin1 +\newcommand{\isasymthreesuperior}{\isamath{\maththreesuperior}} %requires latin1 +\newcommand{\isasymthreequarters}{\isatext{\rm\textthreequarters}} %requires latin1 +\newcommand{\isasymordfeminine}{\isatext{\rm\textordfeminine}} +\newcommand{\isasymordmasculine}{\isatext{\rm\textordmasculine}} +\newcommand{\isasymsection}{\isatext{\rm\S}} +\newcommand{\isasymparagraph}{\isatext{\rm\P}} +\newcommand{\isasymexclamdown}{\isatext{\rm\textexclamdown}} +\newcommand{\isasymquestiondown}{\isatext{\rm\textquestiondown}} +\newcommand{\isasymeuro}{\isatext{\textgreek{\euro}}} %requires greek babel +\newcommand{\isasympounds}{\isamath{\pounds}} +\newcommand{\isasymyen}{\isatext{\yen}} %requires amssymb +\newcommand{\isasymcent}{\isatext{\textcent}} %requires textcomp +\newcommand{\isasymcurrency}{\isatext{\textcurrency}} %requires textcomp +\newcommand{\isasymdegree}{\isatext{\rm\textdegree}} %requires latin1 +\newcommand{\isasymamalg}{\isamath{\amalg}} +\newcommand{\isasymmho}{\isamath{\mho}} %requires amssym +\newcommand{\isasymlozenge}{\isamath{\lozenge}} %requires amssym +\newcommand{\isasymJoin}{\isamath{\Join}} %requires amssym +\newcommand{\isasymwp}{\isamath{\wp}} +\newcommand{\isasymwrong}{\isamath{\wr}} +\newcommand{\isasymstruct}{\isamath{\diamond}} +\newcommand{\isasymacute}{\isatext{\'\relax}} +\newcommand{\isasymindex}{\isatext{\i}} +\newcommand{\isasymdieresis}{\isatext{\"\relax}} +\newcommand{\isasymcedilla}{\isatext{\c\relax}} +\newcommand{\isasymhungarumlaut}{\isatext{\H\relax}} +\newcommand{\isasymspacespace}{\isamath{~~}} diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/llncs.cls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/llncs.cls Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,1189 @@ +% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002) +% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science +% +%% +%% \CharacterTable +%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z +%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z +%% Digits \0\1\2\3\4\5\6\7\8\9 +%% Exclamation \! Double quote \" Hash (number) \# +%% Dollar \$ Percent \% Ampersand \& +%% Acute accent \' Left paren \( Right paren \) +%% Asterisk \* Plus \+ Comma \, +%% Minus \- Point \. Solidus \/ +%% Colon \: Semicolon \; Less than \< +%% Equals \= Greater than \> Question mark \? +%% Commercial at \@ Left bracket \[ Backslash \\ +%% Right bracket \] Circumflex \^ Underscore \_ +%% Grave accent \` Left brace \{ Vertical bar \| +%% Right brace \} Tilde \~} +%% +\NeedsTeXFormat{LaTeX2e}[1995/12/01] +\ProvidesClass{llncs}[2002/01/28 v2.13 +^^J LaTeX document class for Lecture Notes in Computer Science] +% Options +\let\if@envcntreset\iffalse +\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue} +\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y} +\DeclareOption{oribibl}{\let\oribibl=Y} +\let\if@custvec\iftrue +\DeclareOption{orivec}{\let\if@custvec\iffalse} +\let\if@envcntsame\iffalse +\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue} +\let\if@envcntsect\iffalse +\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue} +\let\if@runhead\iffalse +\DeclareOption{runningheads}{\let\if@runhead\iftrue} + +\let\if@openbib\iffalse +\DeclareOption{openbib}{\let\if@openbib\iftrue} + +% languages +\let\switcht@@therlang\relax +\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}} +\def\ds@francais{\def\switcht@@therlang{\switcht@francais}} + +\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}} + +\ProcessOptions + +\LoadClass[twoside]{article} +\RequirePackage{multicol} % needed for the list of participants, index + +\setlength{\textwidth}{12.2cm} +\setlength{\textheight}{19.3cm} +\renewcommand\@pnumwidth{2em} +\renewcommand\@tocrmarg{3.5em} +% +\def\@dottedtocline#1#2#3#4#5{% + \ifnum #1>\c@tocdepth \else + \vskip \z@ \@plus.2\p@ + {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm + \parfillskip -\rightskip \pretolerance=10000 + \parindent #2\relax\@afterindenttrue + \interlinepenalty\@M + \leavevmode + \@tempdima #3\relax + \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip + {#4}\nobreak + \leaders\hbox{$\m@th + \mkern \@dotsep mu\hbox{.}\mkern \@dotsep + mu$}\hfill + \nobreak + \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}% + \par}% + \fi} +% +\def\switcht@albion{% +\def\abstractname{Abstract.} +\def\ackname{Acknowledgement.} +\def\andname{and} +\def\lastandname{\unskip, 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}} +\switcht@albion +% Names of theorem like environments are already defined +% but must be translated if another language is chosen +% +% French section +\def\switcht@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{Preuve}% + \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\switcht@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}% +} + +% Ragged bottom for the actual page +\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil +\global\let\@textbottom\relax}} + +\renewcommand\small{% + \@setfontsize\small\@ixpt{11}% + \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@ + \abovedisplayshortskip \z@ \@plus2\p@ + \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@ + \def\@listi{\leftmargin\leftmargini + \parsep 0\p@ \@plus1\p@ \@minus\p@ + \topsep 8\p@ \@plus2\p@ \@minus4\p@ + \itemsep0\p@}% + \belowdisplayskip \abovedisplayskip +} + +\frenchspacing +\widowpenalty=10000 +\clubpenalty=10000 + +\setlength\oddsidemargin {63\p@} +\setlength\evensidemargin {63\p@} +\setlength\marginparwidth {90\p@} + +\setlength\headsep {16\p@} + +\setlength\footnotesep{7.7\p@} +\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@} +\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@} + +\setcounter{secnumdepth}{2} + +\newcounter {chapter} +\renewcommand\thechapter {\@arabic\c@chapter} + +\newif\if@mainmatter \@mainmattertrue +\newcommand\frontmatter{\cleardoublepage + \@mainmatterfalse\pagenumbering{Roman}} +\newcommand\mainmatter{\cleardoublepage + \@mainmattertrue\pagenumbering{arabic}} +\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi + \@mainmatterfalse} + +\renewcommand\part{\cleardoublepage + \thispagestyle{empty}% + \if@twocolumn + \onecolumn + \@tempswatrue + \else + \@tempswafalse + \fi + \null\vfil + \secdef\@part\@spart} + +\def\@part[#1]#2{% + \ifnum \c@secnumdepth >-2\relax + \refstepcounter{part}% + \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}% + \else + \addcontentsline{toc}{part}{#1}% + \fi + \markboth{}{}% + {\centering + \interlinepenalty \@M + \normalfont + \ifnum \c@secnumdepth >-2\relax + \huge\bfseries \partname~\thepart + \par + \vskip 20\p@ + \fi + \Huge \bfseries #2\par}% + \@endpart} +\def\@spart#1{% + {\centering + \interlinepenalty \@M + \normalfont + \Huge \bfseries #1\par}% + \@endpart} +\def\@endpart{\vfil\newpage + \if@twoside + \null + \thispagestyle{empty}% + \newpage + \fi + \if@tempswa + \twocolumn + \fi} + +\newcommand\chapter{\clearpage + \thispagestyle{empty}% + \global\@topnum\z@ + \@afterindentfalse + \secdef\@chapter\@schapter} +\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne + \if@mainmatter + \refstepcounter{chapter}% + \typeout{\@chapapp\space\thechapter.}% + \addcontentsline{toc}{chapter}% + {\protect\numberline{\thechapter}#1}% + \else + \addcontentsline{toc}{chapter}{#1}% + \fi + \else + \addcontentsline{toc}{chapter}{#1}% + \fi + \chaptermark{#1}% + \addtocontents{lof}{\protect\addvspace{10\p@}}% + \addtocontents{lot}{\protect\addvspace{10\p@}}% + \if@twocolumn + \@topnewpage[\@makechapterhead{#2}]% + \else + \@makechapterhead{#2}% + \@afterheading + \fi} +\def\@makechapterhead#1{% +% \vspace*{50\p@}% + {\centering + \ifnum \c@secnumdepth >\m@ne + \if@mainmatter + \large\bfseries \@chapapp{} \thechapter + \par\nobreak + \vskip 20\p@ + \fi + \fi + \interlinepenalty\@M + \Large \bfseries #1\par\nobreak + \vskip 40\p@ + }} +\def\@schapter#1{\if@twocolumn + \@topnewpage[\@makeschapterhead{#1}]% + \else + \@makeschapterhead{#1}% + \@afterheading + \fi} +\def\@makeschapterhead#1{% +% \vspace*{50\p@}% + {\centering + \normalfont + \interlinepenalty\@M + \Large \bfseries #1\par\nobreak + \vskip 40\p@ + }} + +\renewcommand\section{\@startsection{section}{1}{\z@}% + {-18\p@ \@plus -4\p@ \@minus -4\p@}% + {12\p@ \@plus 4\p@ \@minus 4\p@}% + {\normalfont\large\bfseries\boldmath + \rightskip=\z@ \@plus 8em\pretolerance=10000 }} +\renewcommand\subsection{\@startsection{subsection}{2}{\z@}% + {-18\p@ \@plus -4\p@ \@minus -4\p@}% + {8\p@ \@plus 4\p@ \@minus 4\p@}% + {\normalfont\normalsize\bfseries\boldmath + \rightskip=\z@ \@plus 8em\pretolerance=10000 }} +\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}% + {-18\p@ \@plus -4\p@ \@minus -4\p@}% + {-0.5em \@plus -0.22em \@minus -0.1em}% + {\normalfont\normalsize\bfseries\boldmath}} +\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}% + {-12\p@ \@plus -4\p@ \@minus -4\p@}% + {-0.5em \@plus -0.22em \@minus -0.1em}% + {\normalfont\normalsize\itshape}} +\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use + \string\subparagraph\space with this class}\vskip0.5cm +You should not use \verb|\subparagraph| with this class.\vskip0.5cm} + +\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00} +\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01} +\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02} +\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03} +\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04} +\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05} +\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06} +\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07} +\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08} +\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09} +\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A} + +\let\footnotesize\small + +\if@custvec +\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}} +{\mbox{\boldmath$\textstyle#1$}} +{\mbox{\boldmath$\scriptstyle#1$}} +{\mbox{\boldmath$\scriptscriptstyle#1$}}} +\fi + +\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}} +\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil +\penalty50\hskip1em\null\nobreak\hfil\squareforqed +\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi} + +\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip +\halign{\hfil +$\displaystyle##$\hfil\cr\gets\cr\to\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets +\cr\to\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets +\cr\to\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr +\gets\cr\to\cr}}}}} +\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil +$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr +\noalign{\vskip1.2pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr +\noalign{\vskip1pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr +<\cr +\noalign{\vskip0.9pt}=\cr}}}}} +\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil +$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr +\noalign{\vskip1.2pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr +\noalign{\vskip1pt}=\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr +>\cr +\noalign{\vskip0.9pt}=\cr}}}}} +\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip +\halign{\hfil +$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr +>\cr\noalign{\vskip-1pt}<\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr +>\cr\noalign{\vskip-0.8pt}<\cr}}} +{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr +>\cr\noalign{\vskip-0.3pt}<\cr}}}}} +\def\bbbr{{\rm I\!R}} %reelle Zahlen +\def\bbbm{{\rm I\!M}} +\def\bbbn{{\rm I\!N}} %natuerliche Zahlen +\def\bbbf{{\rm I\!F}} +\def\bbbh{{\rm I\!H}} +\def\bbbk{{\rm I\!K}} +\def\bbbp{{\rm I\!P}} +\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l} +{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}} +\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox +to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox +to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox +to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox +to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}} +\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm +Q$}\hbox{\raise +0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}} +{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise +0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise +0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise +0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}} +\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm +T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox +to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox +to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox +to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}} +\def\bbbs{{\mathchoice +{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox +to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox +to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}} +{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox +to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox +to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox +to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox +to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}} +{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox +to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox +to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}} +\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}} +{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}} +{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}} +{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}} + +\let\ts\, + +\setlength\leftmargini {17\p@} +\setlength\leftmargin {\leftmargini} +\setlength\leftmarginii {\leftmargini} +\setlength\leftmarginiii {\leftmargini} +\setlength\leftmarginiv {\leftmargini} +\setlength \labelsep {.5em} +\setlength \labelwidth{\leftmargini} +\addtolength\labelwidth{-\labelsep} + +\def\@listI{\leftmargin\leftmargini + \parsep 0\p@ \@plus1\p@ \@minus\p@ + \topsep 8\p@ \@plus2\p@ \@minus4\p@ + \itemsep0\p@} +\let\@listi\@listI +\@listi +\def\@listii {\leftmargin\leftmarginii + \labelwidth\leftmarginii + \advance\labelwidth-\labelsep + \topsep 0\p@ \@plus2\p@ \@minus\p@} +\def\@listiii{\leftmargin\leftmarginiii + \labelwidth\leftmarginiii + \advance\labelwidth-\labelsep + \topsep 0\p@ \@plus\p@\@minus\p@ + \parsep \z@ + \partopsep \p@ \@plus\z@ \@minus\p@} + +\renewcommand\labelitemi{\normalfont\bfseries --} +\renewcommand\labelitemii{$\m@th\bullet$} + +\setlength\arraycolsep{1.4\p@} +\setlength\tabcolsep{1.4\p@} + +\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}% + {{\contentsname}}} + \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}} + \def\lastand{\ifnum\value{auco}=2\relax + \unskip{} \andname\ + \else + \unskip \lastandname\ + \fi}% + \def\and{\stepcounter{@auth}\relax + \ifnum\value{@auth}=\value{auco}% + \lastand + \else + \unskip, + \fi}% + \@starttoc{toc}\if@restonecol\twocolumn\fi} + +\def\l@part#1#2{\addpenalty{\@secpenalty}% + \addvspace{2em plus\p@}% % space above part line + \begingroup + \parindent \z@ + \rightskip \z@ plus 5em + \hrule\vskip5pt + \large % same size as for a contribution heading + \bfseries\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\hyperhrefextend{\ifx\hyper@anchor\@undefined\else +{chapter.\thechapter}\fi} + +\def\addnumcontentsmark#1#2#3{% +\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline + {\thechapter}#3}{\thepage}\hyperhrefextend}} +\def\addcontentsmark#1#2#3{% +\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}} +\def\addcontentsmarkwop#1#2#3{% +\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}} + +\def\@adcmk[#1]{\ifcase #1 \or +\def\@gtempa{\addnumcontentsmark}% + \or \def\@gtempa{\addcontentsmark}% + \or \def\@gtempa{\addcontentsmarkwop}% + \fi\@gtempa{toc}{chapter}} +\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}} + +\def\l@chapter#1#2{\addpenalty{-\@highpenalty} + \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup + \parindent \z@ \rightskip \@tocrmarg + \advance\rightskip by 0pt plus 2cm + \parfillskip -\rightskip \pretolerance=10000 + \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip + {\large\bfseries\boldmath#1}\ifx0#2\hfil\null + \else + \nobreak + \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern + \@dotsep mu$}\hfill + \nobreak\hbox to\@pnumwidth{\hss #2}% + \fi\par + \penalty\@highpenalty \endgroup} + +\def\l@title#1#2{\addpenalty{-\@highpenalty} + \addvspace{8pt plus 1pt} + \@tempdima \z@ + \begingroup + \parindent \z@ \rightskip \@tocrmarg + \advance\rightskip by 0pt plus 2cm + \parfillskip -\rightskip \pretolerance=10000 + \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip + #1\nobreak + \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern + \@dotsep mu$}\hfill + \nobreak\hbox to\@pnumwidth{\hss #2}\par + \penalty\@highpenalty \endgroup} + +\def\l@author#1#2{\addpenalty{\@highpenalty} + \@tempdima=\z@ %15\p@ + \begingroup + \parindent \z@ \rightskip \@tocrmarg + \advance\rightskip by 0pt plus 2cm + \pretolerance=10000 + \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip + \textit{#1}\par + \penalty\@highpenalty \endgroup} + +\setcounter{tocdepth}{0} +\newdimen\tocchpnum +\newdimen\tocsecnum +\newdimen\tocsectotal +\newdimen\tocsubsecnum +\newdimen\tocsubsectotal +\newdimen\tocsubsubsecnum +\newdimen\tocsubsubsectotal +\newdimen\tocparanum +\newdimen\tocparatotal +\newdimen\tocsubparanum +\tocchpnum=\z@ % no chapter numbers +\tocsecnum=15\p@ % section 88. plus 2.222pt +\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt +\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt +\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt +\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt +\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\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}} + \@starttoc{lof}\if@restonecol\twocolumn\fi} +\def\l@figure{\@dottedtocline{1}{0em}{1.5em}} + +\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn + \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}} + \@starttoc{lot}\if@restonecol\twocolumn\fi} +\let\l@table\l@figure + +\renewcommand\listoffigures{% + \section*{\listfigurename + \@mkboth{\listfigurename}{\listfigurename}}% + \@starttoc{lof}% + } + +\renewcommand\listoftables{% + \section*{\listtablename + \@mkboth{\listtablename}{\listtablename}}% + \@starttoc{lot}% + } + +\ifx\oribibl\undefined +\ifx\citeauthoryear\undefined +\renewenvironment{thebibliography}[1] + {\section*{\refname} + \def\@biblabel##1{##1.} + \small + \list{\@biblabel{\@arabic\c@enumiv}}% + {\settowidth\labelwidth{\@biblabel{#1}}% + \leftmargin\labelwidth + \advance\leftmargin\labelsep + \if@openbib + \advance\leftmargin\bibindent + \itemindent -\bibindent + \listparindent \itemindent + \parsep \z@ + \fi + \usecounter{enumiv}% + \let\p@enumiv\@empty + \renewcommand\theenumiv{\@arabic\c@enumiv}}% + \if@openbib + \renewcommand\newblock{\par}% + \else + \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}% + \fi + \sloppy\clubpenalty4000\widowpenalty4000% + \sfcode`\.=\@m} + {\def\@noitemerr + {\@latex@warning{Empty `thebibliography' environment}}% + \endlist} +\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw + {\let\protect\noexpand\immediate + \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces} +\newcount\@tempcntc +\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi + \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do + {\@ifundefined + {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries + ?}\@warning + {Citation `\@citeb' on page \thepage \space undefined}}% + {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}% + \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne + \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}% + \else + \advance\@tempcntb\@ne + \ifnum\@tempcntb=\@tempcntc + \else\advance\@tempcntb\m@ne\@citeo + \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}} +\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else + \@citea\def\@citea{,\,\hskip\z@skip}% + \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else + {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else + \def\@citea{--}\fi + \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi} +\else +\renewenvironment{thebibliography}[1] + {\section*{\refname} + \small + \list{}% + {\settowidth\labelwidth{}% + \leftmargin\parindent + \itemindent=-\parindent + \labelsep=\z@ + \if@openbib + \advance\leftmargin\bibindent + \itemindent -\bibindent + \listparindent \itemindent + \parsep \z@ + \fi + \usecounter{enumiv}% + \let\p@enumiv\@empty + \renewcommand\theenumiv{}}% + \if@openbib + \renewcommand\newblock{\par}% + \else + \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}% + \fi + \sloppy\clubpenalty4000\widowpenalty4000% + \sfcode`\.=\@m} + {\def\@noitemerr + {\@latex@warning{Empty `thebibliography' environment}}% + \endlist} + \def\@cite#1{#1}% + \def\@lbibitem[#1]#2{\item[]\if@filesw + {\def\protect##1{\string ##1\space}\immediate + \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces} + \fi +\else +\@cons\@openbib@code{\noexpand\small} +\fi + +\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\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax} + +\renewenvironment{theindex} + {\@mkboth{\indexname}{\indexname}% + \thispagestyle{empty}\parindent\z@ + \parskip\z@ \@plus .3\p@\relax + \let\item\par + \def\,{\relax\ifmmode\mskip\thinmuskip + \else\hskip0.2em\ignorespaces\fi}% + \normalfont\small + \begin{multicols}{2}[\@makeschapterhead{\indexname}]% + } + {\end{multicols}} + +\renewcommand\footnoterule{% + \kern-3\p@ + \hrule\@width 2truecm + \kern2.6\p@} + \newdimen\fnindent + \fnindent1em +\long\def\@makefntext#1{% + \parindent \fnindent% + \leftskip \fnindent% + \noindent + \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1} + +\long\def\@makecaption#1#2{% + \vskip\abovecaptionskip + \sbox\@tempboxa{{\bfseries #1.} #2}% + \ifdim \wd\@tempboxa >\hsize + {\bfseries #1.} #2\par + \else + \global \@minipagefalse + \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}% + \fi + \vskip\belowcaptionskip} + +\def\fps@figure{htbp} +\def\fnum@figure{\figurename\thinspace\thefigure} +\def \@floatboxreset {% + \reset@font + \small + \@setnobreak + \@setminipage +} +\def\fps@table{htbp} +\def\fnum@table{\tablename~\thetable} +\renewenvironment{table} + {\setlength\abovecaptionskip{0\p@}% + \setlength\belowcaptionskip{10\p@}% + \@float{table}} + {\end@float} +\renewenvironment{table*} + {\setlength\abovecaptionskip{0\p@}% + \setlength\belowcaptionskip{10\p@}% + \@dblfloat{table}} + {\end@dblfloat} + +\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} + +% LaTeX does not provide a command to enter the authors institute +% addresses. The \institute command is defined here. + +\newcounter{@inst} +\newcounter{@auth} +\newcounter{auco} +\newdimen\instindent +\newbox\authrun +\newtoks\authorrunning +\newtoks\tocauthor +\newbox\titrun +\newtoks\titlerunning +\newtoks\toctitle + +\def\clearheadinfo{\gdef\@author{No Author Given}% + \gdef\@title{No Title Given}% + \gdef\@subtitle{}% + \gdef\@institute{No Institute Given}% + \gdef\@thanks{}% + \global\titlerunning={}\global\authorrunning={}% + \global\toctitle={}\global\tocauthor={}} + +\def\institute#1{\gdef\@institute{#1}} + +\def\institutename{\par + \begingroup + \parskip=\z@ + \parindent=\z@ + \setcounter{@inst}{1}% + \def\and{\par\stepcounter{@inst}% + \noindent$^{\the@inst}$\enspace\ignorespaces}% + \setbox0=\vbox{\def\thanks##1{}\@institute}% + \ifnum\c@@inst=1\relax + \gdef\fnnstart{0}% + \else + \xdef\fnnstart{\c@@inst}% + \setcounter{@inst}{1}% + \noindent$^{\the@inst}$\enspace + \fi + \ignorespaces + \@institute\par + \endgroup} + +\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or + {\star\star\star}\or \dagger\or \ddagger\or + \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger + \or \ddagger\ddagger \else\@ctrerr\fi}} + +\def\inst#1{\unskip$^{#1}$} +\def\fnmsep{\unskip$^,$} +\def\email#1{{\tt#1}} +\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}% +\@ifpackageloaded{babel}{% +\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}% +\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}% +\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}% +}{\switcht@@therlang}% +} +\def\homedir{\~{ }} + +\def\subtitle#1{\gdef\@subtitle{#1}} +\clearheadinfo + +\renewcommand\maketitle{\newpage + \refstepcounter{chapter}% + \stepcounter{section}% + \setcounter{section}{0}% + \setcounter{subsection}{0}% + \setcounter{figure}{0} + \setcounter{table}{0} + \setcounter{equation}{0} + \setcounter{footnote}{0}% + \begingroup + \parindent=\z@ + \renewcommand\thefootnote{\@fnsymbol\c@footnote}% + \if@twocolumn + \ifnum \col@number=\@ne + \@maketitle + \else + \twocolumn[\@maketitle]% + \fi + \else + \newpage + \global\@topnum\z@ % Prevents figures from going at top of page. + \@maketitle + \fi + \thispagestyle{empty}\@thanks +% + \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}% + \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}% + \instindent=\hsize + \advance\instindent by-\headlineindent + \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else + \addcontentsline{toc}{title}{\the\toctitle}\fi + \if@runhead + \if!\the\titlerunning!\else + \edef\@title{\the\titlerunning}% + \fi + \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}% + \ifdim\wd\titrun>\instindent + \typeout{Title too long for running head. Please supply}% + \typeout{a shorter form with \string\titlerunning\space prior to + \string\maketitle}% + \global\setbox\titrun=\hbox{\small\rm + Title Suppressed Due to Excessive Length}% + \fi + \xdef\@title{\copy\titrun}% + \fi +% + \if!\the\tocauthor!\relax + {\def\and{\noexpand\protect\noexpand\and}% + \protected@xdef\toc@uthor{\@author}}% + \else + \def\\{\noexpand\protect\noexpand\newline}% + \protected@xdef\scratch{\the\tocauthor}% + \protected@xdef\toc@uthor{\scratch}% + \fi + \addcontentsline{toc}{author}{\toc@uthor}% + \if@runhead + \if!\the\authorrunning! + \value{@inst}=\value{@auth}% + \setcounter{@auth}{1}% + \else + \edef\@author{\the\authorrunning}% + \fi + \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}% + \ifdim\wd\authrun>\instindent + \typeout{Names of authors too long for running head. Please supply}% + \typeout{a shorter form with \string\authorrunning\space prior to + \string\maketitle}% + \global\setbox\authrun=\hbox{\small\rm + Authors Suppressed Due to Excessive Length}% + \fi + \xdef\@author{\copy\authrun}% + \markboth{\@author}{\@title}% + \fi + \endgroup + \setcounter{footnote}{\fnnstart}% + \clearheadinfo} +% +\def\@maketitle{\newpage + \markboth{}{}% + \def\lastand{\ifnum\value{@inst}=2\relax + \unskip{} \andname\ + \else + \unskip \lastandname\ + \fi}% + \def\and{\stepcounter{@auth}\relax + \ifnum\value{@auth}=\value{@inst}% + \lastand + \else + \unskip, + \fi}% + \begin{center}% + \let\newline\\ + {\Large \bfseries\boldmath + \pretolerance=10000 + \@title \par}\vskip .8cm +\if!\@subtitle!\else {\large \bfseries\boldmath + \vskip -.65cm + \pretolerance=10000 + \@subtitle \par}\vskip .8cm\fi + \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}% + \def\thanks##1{}\@author}% + \global\value{@inst}=\value{@auth}% + \global\value{auco}=\value{@auth}% + \setcounter{@auth}{1}% +{\lineskip .5em +\noindent\ignorespaces +\@author\vskip.35cm} + {\small\institutename} + \end{center}% + } + +% 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. +% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +\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\ }} + +\if@envcntsect + \def\@thmcountersep{.} + \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape} +\else + \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 % alle Umgebungen wie Theorem. + \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}} +\else % alle Umgebungen mit eigenem Zaehler + \if@envcntsect % mit section numeriert + \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}} + \else % nicht mit section numeriert + \if@envcntreset + \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4} + \@addtoreset{#1}{section}} + \else + \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4} + \@addtoreset{#1}{chapter}}% + \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} +\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily} +\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape} +\spn@wtheorem{note}{Note}{\itshape}{\rmfamily} +\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily} +\spn@wtheorem{property}{Property}{\itshape}{\rmfamily} +\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape} +\spn@wtheorem{question}{Question}{\itshape}{\rmfamily} +\spn@wtheorem{solution}{Solution}{\itshape}{\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\ }} + } + +\renewenvironment{abstract}{% + \list{}{\advance\topsep by0.35cm\relax\small + \leftmargin=1cm + \labelwidth=\z@ + \listparindent=\z@ + \itemindent\listparindent + \rightmargin\leftmargin}\item[\hskip\labelsep + \bfseries\abstractname]} + {\endlist} + +\newdimen\headlineindent % dimension for space between +\headlineindent=1.166cm % number and text of headings. + +\def\ps@headings{\let\@mkboth\@gobbletwo + \let\@oddfoot\@empty\let\@evenfoot\@empty + \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}% + \leftmark\hfil} + \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}% + \llap{\thepage}} + \def\chaptermark##1{}% + \def\sectionmark##1{}% + \def\subsectionmark##1{}} + +\def\ps@titlepage{\let\@mkboth\@gobbletwo + \let\@oddfoot\@empty\let\@evenfoot\@empty + \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}% + \hfil} + \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}% + \llap{\thepage}} + \def\chaptermark##1{}% + \def\sectionmark##1{}% + \def\subsectionmark##1{}} + +\if@runhead\ps@headings\else +\ps@empty\fi + +\setlength\arraycolsep{1.4\p@} +\setlength\tabcolsep{1.4\p@} + +\endinput +%end of file llncs.cls diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/pdfsetup.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/pdfsetup.sty Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,12 @@ +%% +%% Author: Markus Wenzel, TU Muenchen +%% License: GPL (GNU GENERAL PUBLIC LICENSE) +%% +%% smart url or hyperref setup +%% + +\@ifundefined{pdfoutput} +{\usepackage{url}} +{\usepackage{color}\definecolor{darkblue}{rgb}{0,0,0.5} + \usepackage[pdftex,a4paper,colorlinks=true,linkcolor=darkblue,citecolor=darkblue,filecolor=darkblue,pagecolor=darkblue,urlcolor=darkblue]{hyperref} + \IfFileExists{thumbpdf.sty}{\usepackage{thumbpdf}}{}} diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/root.bib --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/root.bib Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,52 @@ +@string{LNCS="Lect.\ Notes in Comp.\ Sci."} +@string{Springer="Springer-Verlag"} +@string{JAR="J. Automated Reasoning"} + +@InProceedings{BauerW-TPHOLs01, +author={Gertrud Bauer and Markus Wenzel}, +title={Calculational Reasoning Revisited --- an {Isabelle/Isar} Experience}, +booktitle={Theorem Proving in Higher Order Logics, TPHOLs 2001}, +editor={R. Boulton and P. Jackson}, +year=2001,publisher=Springer,series=LNCS,volume=2152,pages="75--90"} + +@book{LCF,author="M.C.J. Gordon and Robin Milner and C.P. Wadsworth", +title="Edinburgh {LCF}: a Mechanised Logic of Computation", +publisher=Springer,series=LNCS,volume=78,year=1979} + +@InProceedings{KWP-TPHOLs99, +author={Florian Kamm{\"u}ller and Markus Wenzel and Lawrence C. Paulson}, +title={Locales: A Sectioning Concept for {Isabelle}}, +booktitle={Theorem Proving in Higher Order Logics, TPHOLs'99}, +editor={Y. Bertot and G. Dowek and A. Hirschowitz and C. Paulin and L. Thery}, +year=1999,publisher=Springer,series=LNCS,volume=1690,pages="149--165"} + +@book{LNCS2283,author={Tobias Nipkow and Lawrence Paulson and Markus Wenzel}, +title="Isabelle/HOL --- A Proof Assistant for Higher-Order Logic", +publisher=Springer,series=LNCS,volume=2283,year=2002, +note={\url{http://www.in.tum.de/~nipkow/LNCS2283/}}} + +@article{KleinN-TCS,author={Gerwin Klein and Tobias Nipkow}, +title={Verified Bytecode Verifiers}, +journal=TCS,year=2002,note={To appear. +\url{http://www.in.tum.de/~nipkow/pubs/tcs03.html}}} + +@InProceedings{Rudnicki92,author={P. Rudnicki}, +title={An Overview of the {Mizar} Project}, +booktitle={Workshop on Types for Proofs and Programs}, +year=1992,organization={Chalmers University of Technology}} + +@manual{Isar-Ref-Man,author="Markus Wenzel", +title="The Isabelle/Isar Reference Manual", +organization={Technische Universit{\"a}t M{\"u}nchen},year=2002, +note={\url{http://isabelle.in.tum.de/dist/Isabelle2002/doc/isar-ref.pdf}}} + +@phdthesis{Wenzel-PhD,author={Markus Wenzel},title={Isabelle/Isar --- A +Versatile Environment for Human-Readable Formal Proof Documents}, +school={Institut f{\"u}r Informatik, Technische Universit{\"a}t M{\"u}nchen}, +year=2002, +note={\url{http://tumb1.biblio.tu-muenchen.de/publ/diss/in/2002/wenzel.html}}} + +@article{WenzelW-JAR,author={Markus Wenzel and Freek Wiedijk}, +title={A comparison of the mathematical proof languages {Mizar} and {Isar}}, +journal=JAR,year=2003,note={To appear}} + diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/root.tex Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,50 @@ +\documentclass[envcountsame]{llncs} +%\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym,pdfsetup} + +%for best-style documents ... +\urlstyle{rm} +%\isabellestyle{it} + +\newcommand{\tweakskip}{\vspace{-\medskipamount}} +\newcommand{\Tweakskip}{\tweakskip\tweakskip} + +\pagestyle{plain} + +\begin{document} + +\title{%A Compact Introduction to +Structured Proofs in Isar/HOL} +\author{Tobias Nipkow} +\institute{Institut f{\"u}r Informatik, TU M{\"u}nchen\\ + {\small\url{http://www.in.tum.de/~nipkow/}}} +\date{} +\maketitle + +\begin{abstract} + Isar is an extension of the theorem prover Isabelle with a language + for writing human-readable structured proofs. This paper is an + introduction to the basic constructs of this language. +% It is aimed at potential users of Isar +% but also discusses the design rationals +% behind the language and its constructs. +\end{abstract} + +\input{intro.tex} +\input{Logic.tex} +\Tweakskip\Tweakskip +\input{Induction.tex} + +%\Tweakskip +\small +\paragraph{Acknowledgement} +I am deeply indebted to Markus Wenzel for conceiving Isar. Clemens Ballarin, +Gertrud Bauer, Stefan Berghofer, Gerwin Klein, Norbert Schirmer, +Markus Wenzel and Freek Wiedijk commented on and improved this paper. + +\begingroup +\bibliographystyle{plain} \small\raggedright\frenchspacing +\bibliography{root} +\endgroup + +\end{document} diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/document/session.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/document/session.tex Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,8 @@ +\input{Logic.tex} + +\input{Induction.tex} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 75a399c2781f -r 454a2ad0c381 doc-src/IsarOverview/Isar/makeDemo --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarOverview/Isar/makeDemo Mon May 12 11:33:55 2003 +0200 @@ -0,0 +1,36 @@ +#!/usr/bin/perl -w + +sub doit { + my ($file) = @_; + + open (FILE, $file) || die $!; + undef $/; $text = ; $/ = "\n"; + close FILE || die $!; + + $_ = $text; + + s/text_raw\{\*([^*]|\*[^}])*\*\}//sg; # actual work done here + s/text\{\*([^*]|\*[^}])*\*\}//sg; # actual work done here + s/\(\*<\*\)//sg; + s/\(\*>\*\)//sg; + s/--(\ )*"([^"])*"//sg; + s/--(\ )*\{\*([^*]|\*[^}])*\*\}//sg; + + $result = $_; + + if ($text ne $result) { + print STDERR "fixing $file\n"; +# if (! -f "$file~~") { +# rename $file, "$file~~" || die $!; +# } + open (FILE, "> Demo/$file") || die $!; + print FILE $result; + close FILE || die $!; + } +} + + +foreach $file (@ARGV) { + eval { &doit($file); }; + if ($@) { print STDERR "*** doit $file: ", $@, "\n"; } +} diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/IsaMakefile --- a/doc-src/TutorialI/IsarOverview/IsaMakefile Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,26 +0,0 @@ -## targets - -default: Isar - -## global settings - -SRC = $(ISABELLE_HOME)/src -OUT = $(ISABELLE_OUTPUT) -LOG = $(OUT)/log -USEDIR = $(ISATOOL) usedir -i true -d ps -D document -v true - - -## Isar - -Isar: $(LOG)/HOL-Isar.gz - -$(LOG)/HOL-Isar.gz: Isar/ROOT.ML Isar/document/intro.tex \ - Isar/document/root.tex Isar/document/root.bib Isar/*.thy - @$(USEDIR) HOL Isar - - -## clean - -clean: - @rm -f $(LOG)/HOL-Isar.gz - diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/Calc.thy --- a/doc-src/TutorialI/IsarOverview/Isar/Calc.thy Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,49 +0,0 @@ -theory Calc = Main: - -subsection{* Chains of equalities *} - -axclass - group < zero, plus, minus - assoc: "(x + y) + z = x + (y + z)" - left_0: "0 + x = x" - left_minus: "-x + x = 0" - -theorem right_minus: "x + -x = (0::'a::group)" -proof - - have "x + -x = (-(-x) + -x) + (x + -x)" - by (simp only: left_0 left_minus) - also have "... = -(-x) + ((-x + x) + -x)" - by (simp only: group.assoc) - also have "... = 0" - by (simp only: left_0 left_minus) - finally show ?thesis . -qed - -subsection{* Inequalities and substitution *} - -lemmas distrib = zadd_zmult_distrib zadd_zmult_distrib2 - zdiff_zmult_distrib zdiff_zmult_distrib2 -declare numeral_2_eq_2[simp] - - -lemma fixes a :: int shows "(a+b)\ \ 2*(a\ + b\)" -proof - - have "(a+b)\ \ (a+b)\ + (a-b)\" by simp - also have "(a+b)\ \ a\ + b\ + 2*a*b" by(simp add:distrib) - also have "(a-b)\ = a\ + b\ - 2*a*b" by(simp add:distrib) - finally show ?thesis by simp -qed - - -subsection{* More transitivity *} - -lemma assumes R: "(a,b) \ R" "(b,c) \ R" "(c,d) \ R" - shows "(a,d) \ R\<^sup>*" -proof - - have "(a,b) \ R\<^sup>*" .. - also have "(b,c) \ R\<^sup>*" .. - also have "(c,d) \ R\<^sup>*" .. - finally show ?thesis . -qed - -end \ No newline at end of file diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/Induction.thy --- a/doc-src/TutorialI/IsarOverview/Isar/Induction.thy Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,293 +0,0 @@ -(*<*)theory Induction = Main:(*>*) - -section{*Case distinction and induction \label{sec:Induct}*} - -text{* Computer science applications abound with inductively defined -structures, which is why we treat them in more detail. HOL already -comes with a datatype of lists with the two constructors @{text Nil} -and @{text Cons}. @{text Nil} is written @{term"[]"} and @{text"Cons x -xs"} is written @{term"x # xs"}. *} - -subsection{*Case distinction\label{sec:CaseDistinction}*} - -text{* We have already met the @{text cases} method for performing -binary case splits. Here is another example: *} -lemma "\ A \ A" -proof cases - assume "A" thus ?thesis .. -next - assume "\ A" thus ?thesis .. -qed - -text{*\noindent The two cases must come in this order because @{text -cases} merely abbreviates @{text"(rule case_split_thm)"} where -@{thm[source] case_split_thm} is @{thm case_split_thm}. If we reverse -the order of the two cases in the proof, the first case would prove -@{prop"\ A \ \ A \ A"} which would solve the first premise of -@{thm[source] case_split_thm}, instantiating @{text ?P} with @{term "\ -A"}, thus making the second premise @{prop"\ \ A \ \ A \ A"}. -Therefore the order of subgoals is not always completely arbitrary. - -The above proof is appropriate if @{term A} is textually small. -However, if @{term A} is large, we do not want to repeat it. This can -be avoided by the following idiom *} - -lemma "\ A \ A" -proof (cases "A") - case True thus ?thesis .. -next - case False thus ?thesis .. -qed - -text{*\noindent which is like the previous proof but instantiates -@{text ?P} right away with @{term A}. Thus we could prove the two -cases in any order. The phrase `\isakeyword{case}~@{text True}' -abbreviates `\isakeyword{assume}~@{text"True: A"}' and analogously for -@{text"False"} and @{prop"\A"}. - -The same game can be played with other datatypes, for example lists, -where @{term tl} is the tail of a list, and @{text length} returns a -natural number (remember: $0-1=0$): -*} -(*<*)declare length_tl[simp del](*>*) -lemma "length(tl xs) = length xs - 1" -proof (cases xs) - case Nil thus ?thesis by simp -next - case Cons thus ?thesis by simp -qed -text{*\noindent Here `\isakeyword{case}~@{text Nil}' abbreviates -`\isakeyword{assume}~@{text"Nil:"}~@{prop"xs = []"}' and -`\isakeyword{case}~@{text Cons}' -abbreviates `\isakeyword{fix}~@{text"? ??"} -\isakeyword{assume}~@{text"Cons:"}~@{text"xs = ? # ??"}' -where @{text"?"} and @{text"??"} -stand for variable names that have been chosen by the system. -Therefore we cannot refer to them. -Luckily, this proof is simple enough we do not need to refer to them. -However, sometimes one may have to. Hence Isar offers a simple scheme for -naming those variables: replace the anonymous @{text Cons} by -@{text"(Cons y ys)"}, which abbreviates `\isakeyword{fix}~@{text"y ys"} -\isakeyword{assume}~@{text"Cons:"}~@{text"xs = y # ys"}'. -In each \isakeyword{case} the assumption can be -referred to inside the proof by the name of the constructor. In -Section~\ref{sec:full-Ind} below we will come across an example -of this. *} - -subsection{*Structural induction*} - -text{* We start with an inductive proof where both cases are proved automatically: *} -lemma "2 * (\iii"} or @{text"\"}\label{sec:full-Ind}*} - -text{* Let us now consider the situation where the goal to be proved contains -@{text"\"} or @{text"\"}, say @{prop"\x. P x \ Q x"} --- motivation and a -real example follow shortly. This means that in each case of the induction, -@{text ?case} would be of the form @{prop"\x. P' x \ Q' x"}. Thus the -first proof steps will be the canonical ones, fixing @{text x} and assuming -@{prop"P' x"}. To avoid this tedium, induction performs these steps -automatically: for example in case @{text"(Suc n)"}, @{text ?case} is only -@{prop"Q' x"} whereas the assumptions (named @{term Suc}!) contain both the -usual induction hypothesis \emph{and} @{prop"P' x"}. -It should be clear how this generalises to more complex formulae. - -As an example we will now prove complete induction via -structural induction. *} - -lemma assumes A: "(\n. (\m. m < n \ P m) \ P n)" - shows "P(n::nat)" -proof (rule A) - show "\m. m < n \ P m" - proof (induct n) - case 0 thus ?case by simp - next - case (Suc n) -- {*\isakeyword{fix} @{term m} \isakeyword{assume} @{text Suc}: @{text[source]"?m < n \ P ?m"} @{prop[source]"m < Suc n"}*} - show ?case -- {*@{term ?case}*} - proof cases - assume eq: "m = n" - from Suc and A have "P n" by blast - with eq show "P m" by simp - next - assume "m \ n" - with Suc have "m < n" by arith - thus "P m" by(rule Suc) - qed - qed -qed -text{* \noindent Given the explanations above and the comments in the -proof text (only necessary for novices), the proof should be quite -readable. - -The statement of the lemma is interesting because it deviates from the style in -the Tutorial~\cite{LNCS2283}, which suggests to introduce @{text"\"} or -@{text"\"} into a theorem to strengthen it for induction. In Isar -proofs we can use @{text"\"} and @{text"\"} instead. This simplifies the -proof and means we do not have to convert between the two kinds of -connectives. - -Note that in a nested induction over the same data type, the inner -case labels hide the outer ones of the same name. If you want to refer -to the outer ones inside, you need to name them on the outside, e.g.\ -\isakeyword{note}~@{text"outer_IH = Suc"}. *} - -subsection{*Rule induction*} - -text{* HOL also supports inductively defined sets. See \cite{LNCS2283} -for details. As an example we define our own version of the reflexive -transitive closure of a relation --- HOL provides a predefined one as well.*} -consts rtc :: "('a \ 'a)set \ ('a \ 'a)set" ("_*" [1000] 999) -inductive "r*" -intros -refl: "(x,x) \ r*" -step: "\ (x,y) \ r; (y,z) \ r* \ \ (x,z) \ r*" - -text{* \noindent -First the constant is declared as a function on binary -relations (with concrete syntax @{term"r*"} instead of @{text"rtc -r"}), then the defining clauses are given. We will now prove that -@{term"r*"} is indeed transitive: *} - -lemma assumes A: "(x,y) \ r*" shows "(y,z) \ r* \ (x,z) \ r*" -using A -proof induct - case refl thus ?case . -next - case step thus ?case by(blast intro: rtc.step) -qed -text{*\noindent Rule induction is triggered by a fact $(x_1,\dots,x_n) -\in R$ piped into the proof, here \isakeyword{using}~@{text A}. The -proof itself follows the inductive definition very -closely: there is one case for each rule, and it has the same name as -the rule, analogous to structural induction. - -However, this proof is rather terse. Here is a more readable version: -*} - -lemma assumes A: "(x,y) \ r*" and B: "(y,z) \ r*" - shows "(x,z) \ r*" -proof - - from A B show ?thesis - proof induct - fix x assume "(x,z) \ r*" -- {*@{text B}[@{text y} := @{text x}]*} - thus "(x,z) \ r*" . - next - fix x' x y - assume 1: "(x',x) \ r" and - IH: "(y,z) \ r* \ (x,z) \ r*" and - B: "(y,z) \ r*" - from 1 IH[OF B] show "(x',z) \ r*" by(rule rtc.step) - qed -qed -text{*\noindent We start the proof with \isakeyword{from}~@{text"A -B"}. Only @{text A} is ``consumed'' by the induction step. -Since @{text B} is left over we don't just prove @{text -?thesis} but @{text"B \ ?thesis"}, just as in the previous proof. The -base case is trivial. In the assumptions for the induction step we can -see very clearly how things fit together and permit ourselves the -obvious forward step @{text"IH[OF B]"}. - -The notation `\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}' -is also supported for inductive definitions. The \emph{constructor} is (the -name of) the rule and the \emph{vars} fix the free variables in the -rule; the order of the \emph{vars} must correspond to the -\emph{alphabetical order} of the variables as they appear in the rule. -For example, we could start the above detailed proof of the induction -with \isakeyword{case}~\isa{(step x' x y)}. However, we can then only -refer to the assumptions named \isa{step} collectively and not -individually, as the above proof requires. *} - -subsection{*More induction*} - -text{* We close the section by demonstrating how arbitrary induction -rules are applied. As a simple example we have chosen recursion -induction, i.e.\ induction based on a recursive function -definition. However, most of what we show works for induction in -general. - -The example is an unusual definition of rotation: *} - -consts rot :: "'a list \ 'a list" -recdef rot "measure length" --"for the internal termination proof" -"rot [] = []" -"rot [x] = [x]" -"rot (x#y#zs) = y # rot(x#zs)" -text{*\noindent This yields, among other things, the induction rule -@{thm[source]rot.induct}: @{thm[display]rot.induct[no_vars]} -In the following proof we rely on a default naming scheme for cases: they are -called 1, 2, etc, unless they have been named explicitly. The latter happens -only with datatypes and inductively defined sets, but not with recursive -functions. *} - -lemma "xs \ [] \ rot xs = tl xs @ [hd xs]" -proof (induct xs rule: rot.induct) - case 1 thus ?case by simp -next - case 2 show ?case by simp -next - case (3 a b cs) - have "rot (a # b # cs) = b # rot(a # cs)" by simp - also have "\ = b # tl(a # cs) @ [hd(a # cs)]" by(simp add:3) - also have "\ = tl (a # b # cs) @ [hd (a # b # cs)]" by simp - finally show ?case . -qed - -text{*\noindent -The third case is only shown in gory detail (see \cite{BauerW-TPHOLs01} -for how to reason with chains of equations) to demonstrate that the -`\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}' notation also -works for arbitrary induction theorems with numbered cases. The order -of the \emph{vars} corresponds to the order of the -@{text"\"}-quantified variables in each case of the induction -theorem. For induction theorems produced by \isakeyword{recdef} it is -the order in which the variables appear on the left-hand side of the -equation. - -The proof is so simple that it can be condensed to -*} - -(*<*)lemma "xs \ [] \ rot xs = tl xs @ [hd xs]"(*>*) -by (induct xs rule: rot.induct, simp_all) - -(*<*)end(*>*) diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/Logic.thy --- a/doc-src/TutorialI/IsarOverview/Isar/Logic.thy Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,644 +0,0 @@ -(*<*)theory Logic = Main:(*>*) - -section{*Logic \label{sec:Logic}*} - -subsection{*Propositional logic*} - -subsubsection{*Introduction rules*} - -text{* We start with a really trivial toy proof to introduce the basic -features of structured proofs. *} -lemma "A \ A" -proof (rule impI) - assume a: "A" - show "A" by(rule a) -qed -text{*\noindent -The operational reading: the \isakeyword{assume}-\isakeyword{show} -block proves @{prop"A \ A"} (@{term a} is a degenerate rule (no -assumptions) that proves @{term A} outright), which rule -@{thm[source]impI} (@{thm impI}) turns into the desired @{prop"A \ -A"}. However, this text is much too detailed for comfort. Therefore -Isar implements the following principle: \begin{quote}\em Command -\isakeyword{proof} automatically tries to select an introduction rule -based on the goal and a predefined list of rules. \end{quote} Here -@{thm[source]impI} is applied automatically: *} - -lemma "A \ A" -proof - assume a: A - show A by(rule a) -qed - -text{*\noindent Single-identifier formulae such as @{term A} need not -be enclosed in double quotes. However, we will continue to do so for -uniformity. - -Trivial proofs, in particular those by assumption, should be trivial -to perform. Proof ``.'' does just that (and a bit more). Thus -naming of assumptions is often superfluous: *} -lemma "A \ A" -proof - assume "A" - show "A" . -qed - -text{* To hide proofs by assumption further, \isakeyword{by}@{text"(method)"} -first applies @{text method} and then tries to solve all remaining subgoals -by assumption: *} -lemma "A \ A \ A" -proof - assume "A" - show "A \ A" by(rule conjI) -qed -text{*\noindent Rule @{thm[source]conjI} is of course @{thm conjI}. -A drawback of implicit proofs by assumption is that it -is no longer obvious where an assumption is used. - -Proofs of the form \isakeyword{by}@{text"(rule"}~\emph{name}@{text")"} -can be abbreviated to ``..'' if \emph{name} refers to one of the -predefined introduction rules (or elimination rules, see below): *} - -lemma "A \ A \ A" -proof - assume "A" - show "A \ A" .. -qed -text{*\noindent -This is what happens: first the matching introduction rule @{thm[source]conjI} -is applied (first ``.''), then the two subgoals are solved by assumption -(second ``.''). *} - -subsubsection{*Elimination rules*} - -text{*A typical elimination rule is @{thm[source]conjE}, $\land$-elimination: -@{thm[display,indent=5]conjE} In the following proof it is applied -by hand, after its first (\emph{major}) premise has been eliminated via -@{text"[OF AB]"}: *} -lemma "A \ B \ B \ A" -proof - assume AB: "A \ B" - show "B \ A" - proof (rule conjE[OF AB]) -- {*@{text"conjE[OF AB]"}: @{thm conjE[OF AB]} *} - assume "A" "B" - show ?thesis .. - qed -qed -text{*\noindent Note that the term @{text"?thesis"} always stands for the -``current goal'', i.e.\ the enclosing \isakeyword{show} (or -\isakeyword{have}) statement. - -This is too much proof text. Elimination rules should be selected -automatically based on their major premise, the formula or rather connective -to be eliminated. In Isar they are triggered by facts being fed -\emph{into} a proof. Syntax: -\begin{center} -\isakeyword{from} \emph{fact} \isakeyword{show} \emph{proposition} \emph{proof} -\end{center} -where \emph{fact} stands for the name of a previously proved -proposition, e.g.\ an assumption, an intermediate result or some global -theorem, which may also be modified with @{text OF} etc. -The \emph{fact} is ``piped'' into the \emph{proof}, which can deal with it -how it chooses. If the \emph{proof} starts with a plain \isakeyword{proof}, -an elimination rule (from a predefined list) is applied -whose first premise is solved by the \emph{fact}. Thus the proof above -is equivalent to the following one: *} - -lemma "A \ B \ B \ A" -proof - assume AB: "A \ B" - from AB show "B \ A" - proof - assume "A" "B" - show ?thesis .. - qed -qed - -text{* Now we come to a second important principle: -\begin{quote}\em -Try to arrange the sequence of propositions in a UNIX-like pipe, -such that the proof of each proposition builds on the previous proposition. -\end{quote} -The previous proposition can be referred to via the fact @{text this}. -This greatly reduces the need for explicit naming of propositions: -*} -lemma "A \ B \ B \ A" -proof - assume "A \ B" - from this show "B \ A" - proof - assume "A" "B" - show ?thesis .. - qed -qed - -text{*\noindent Because of the frequency of \isakeyword{from}~@{text -this}, Isar provides two abbreviations: -\begin{center} -\begin{tabular}{r@ {\quad=\quad}l} -\isakeyword{then} & \isakeyword{from} @{text this} \\ -\isakeyword{thus} & \isakeyword{then} \isakeyword{show} -\end{tabular} -\end{center} - -Here is an alternative proof that operates purely by forward reasoning: *} -lemma "A \ B \ B \ A" -proof - assume ab: "A \ B" - from ab have a: "A" .. - from ab have b: "B" .. - from b a show "B \ A" .. -qed - -text{*\noindent It is worth examining this text in detail because it -exhibits a number of new concepts. For a start, it is the first time -we have proved intermediate propositions (\isakeyword{have}) on the -way to the final \isakeyword{show}. This is the norm in nontrivial -proofs where one cannot bridge the gap between the assumptions and the -conclusion in one step. To understand how the proof works we need to -explain more Isar details. - -Method @{text rule} can be given a list of rules, in which case -@{text"(rule"}~\textit{rules}@{text")"} applies the first matching -rule in the list \textit{rules}. Command \isakeyword{from} can be -followed by any number of facts. Given \isakeyword{from}~@{text -f}$_1$~\dots~@{text f}$_n$, the proof step -@{text"(rule"}~\textit{rules}@{text")"} following a \isakeyword{have} -or \isakeyword{show} searches \textit{rules} for a rule whose first -$n$ premises can be proved by @{text f}$_1$~\dots~@{text f}$_n$ in the -given order. Finally one needs to know that ``..'' is short for -@{text"by(rule"}~\textit{elim-rules intro-rules}@{text")"} (or -@{text"by(rule"}~\textit{intro-rules}@{text")"} if there are no facts -fed into the proof), i.e.\ elimination rules are tried before -introduction rules. - -Thus in the above proof both \isakeyword{have}s are proved via -@{thm[source]conjE} triggered by \isakeyword{from}~@{text ab} whereas -in the \isakeyword{show} step no elimination rule is applicable and -the proof succeeds with @{thm[source]conjI}. The latter would fail had -we written \isakeyword{from}~@{text"a b"} instead of -\isakeyword{from}~@{text"b a"}. - -Proofs starting with a plain @{text proof} behave the same because the -latter is short for @{text"proof (rule"}~\textit{elim-rules -intro-rules}@{text")"} (or @{text"proof -(rule"}~\textit{intro-rules}@{text")"} if there are no facts fed into -the proof). *} - -subsection{*More constructs*} - -text{* In the previous proof of @{prop"A \ B \ B \ A"} we needed to feed -more than one fact into a proof step, a frequent situation. Then the -UNIX-pipe model appears to break down and we need to name the different -facts to refer to them. But this can be avoided: -*} -lemma "A \ B \ B \ A" -proof - assume ab: "A \ B" - from ab have "B" .. - moreover - from ab have "A" .. - ultimately show "B \ A" .. -qed -text{*\noindent You can combine any number of facts @{term A1} \dots\ @{term -An} into a sequence by separating their proofs with -\isakeyword{moreover}. After the final fact, \isakeyword{ultimately} stands -for \isakeyword{from}~@{term A1}~\dots~@{term An}. This avoids having to -introduce names for all of the sequence elements. *} - -text{* Although we have only seen a few introduction and elimination rules so -far, Isar's predefined rules include all the usual natural deduction -rules. We conclude our exposition of propositional logic with an extended -example --- which rules are used implicitly where? *} -lemma "\ (A \ B) \ \ A \ \ B" -proof - assume n: "\ (A \ B)" - show "\ A \ \ B" - proof (rule ccontr) - assume nn: "\ (\ A \ \ B)" - have "\ A" - proof - assume "A" - have "\ B" - proof - assume "B" - have "A \ B" .. - with n show False .. - qed - hence "\ A \ \ B" .. - with nn show False .. - qed - hence "\ A \ \ B" .. - with nn show False .. - qed -qed -text{*\noindent -Rule @{thm[source]ccontr} (``classical contradiction'') is -@{thm ccontr[no_vars]}. -Apart from demonstrating the strangeness of classical -arguments by contradiction, this example also introduces two new -abbreviations: -\begin{center} -\begin{tabular}{l@ {\quad=\quad}l} -\isakeyword{hence} & \isakeyword{then} \isakeyword{have} \\ -\isakeyword{with}~\emph{facts} & -\isakeyword{from}~\emph{facts} @{text this} -\end{tabular} -\end{center} -*} - -subsection{*Avoiding duplication*} - -text{* So far our examples have been a bit unnatural: normally we want to -prove rules expressed with @{text"\"}, not @{text"\"}. Here is an example: -*} -lemma "A \ B \ B \ A" -proof - assume "A \ B" thus "B" .. -next - assume "A \ B" thus "A" .. -qed -text{*\noindent The \isakeyword{proof} always works on the conclusion, -@{prop"B \ A"} in our case, thus selecting $\land$-introduction. Hence -we must show @{prop B} and @{prop A}; both are proved by -$\land$-elimination and the proofs are separated by \isakeyword{next}: -\begin{description} -\item[\isakeyword{next}] deals with multiple subgoals. For example, -when showing @{term"A \ B"} we need to show both @{term A} and @{term -B}. Each subgoal is proved separately, in \emph{any} order. The -individual proofs are separated by \isakeyword{next}. \footnote{Each -\isakeyword{show} must prove one of the pending subgoals. If a -\isakeyword{show} matches multiple subgoals, e.g.\ if the subgoals -contain ?-variables, the first one is proved. Thus the order in which -the subgoals are proved can matter --- see -\S\ref{sec:CaseDistinction} for an example.} - -Strictly speaking \isakeyword{next} is only required if the subgoals -are proved in different assumption contexts which need to be -separated, which is not the case above. For clarity we -have employed \isakeyword{next} anyway and will continue to do so. -\end{description} - -This is all very well as long as formulae are small. Let us now look at some -devices to avoid repeating (possibly large) formulae. A very general method -is pattern matching: *} - -lemma "large_A \ large_B \ large_B \ large_A" - (is "?AB \ ?B \ ?A") -proof - assume "?AB" thus "?B" .. -next - assume "?AB" thus "?A" .. -qed -text{*\noindent Any formula may be followed by -@{text"("}\isakeyword{is}~\emph{pattern}@{text")"} which causes the pattern -to be matched against the formula, instantiating the @{text"?"}-variables in -the pattern. Subsequent uses of these variables in other terms causes -them to be replaced by the terms they stand for. - -We can simplify things even more by stating the theorem by means of the -\isakeyword{assumes} and \isakeyword{shows} elements which allow direct -naming of assumptions: *} - -lemma assumes AB: "large_A \ large_B" - shows "large_B \ large_A" (is "?B \ ?A") -proof - from AB show "?B" .. -next - from AB show "?A" .. -qed -text{*\noindent Note the difference between @{text ?AB}, a term, and -@{text AB}, a fact. - -Finally we want to start the proof with $\land$-elimination so we -don't have to perform it twice, as above. Here is a slick way to -achieve this: *} - -lemma assumes AB: "large_A \ large_B" - shows "large_B \ large_A" (is "?B \ ?A") -using AB -proof - assume "?A" "?B" show ?thesis .. -qed -text{*\noindent Command \isakeyword{using} can appear before a proof -and adds further facts to those piped into the proof. Here @{text AB} -is the only such fact and it triggers $\land$-elimination. Another -frequent idiom is as follows: -\begin{center} -\isakeyword{from} \emph{major-facts}~ -\isakeyword{show} \emph{proposition}~ -\isakeyword{using} \emph{minor-facts}~ -\emph{proof} -\end{center} - -Sometimes it is necessary to suppress the implicit application of rules in a -\isakeyword{proof}. For example \isakeyword{show}~@{prop"A \ B"} would -trigger $\lor$-introduction, requiring us to prove @{prop A}. A simple -``@{text"-"}'' prevents this \emph{faux pas}: *} - -lemma assumes AB: "A \ B" shows "B \ A" -proof - - from AB show ?thesis - proof - assume A show ?thesis .. - next - assume B show ?thesis .. - qed -qed - - -subsection{*Predicate calculus*} - -text{* Command \isakeyword{fix} introduces new local variables into a -proof. The pair \isakeyword{fix}-\isakeyword{show} corresponds to @{text"\"} -(the universal quantifier at the -meta-level) just like \isakeyword{assume}-\isakeyword{show} corresponds to -@{text"\"}. Here is a sample proof, annotated with the rules that are -applied implicitly: *} -lemma assumes P: "\x. P x" shows "\x. P(f x)" -proof --"@{thm[source]allI}: @{thm allI}" - fix a - from P show "P(f a)" .. --"@{thm[source]allE}: @{thm allE}" -qed -text{*\noindent Note that in the proof we have chosen to call the bound -variable @{term a} instead of @{term x} merely to show that the choice of -local names is irrelevant. - -Next we look at @{text"\"} which is a bit more tricky. -*} - -lemma assumes Pf: "\x. P(f x)" shows "\y. P y" -proof - - from Pf show ?thesis - proof -- "@{thm[source]exE}: @{thm exE}" - fix x - assume "P(f x)" - show ?thesis .. -- "@{thm[source]exI}: @{thm exI}" - qed -qed -text{*\noindent Explicit $\exists$-elimination as seen above can become -cumbersome in practice. The derived Isar language element -\isakeyword{obtain} provides a more appealing form of generalised -existence reasoning: *} - -lemma assumes Pf: "\x. P(f x)" shows "\y. P y" -proof - - from Pf obtain x where "P(f x)" .. - thus "\y. P y" .. -qed -text{*\noindent Note how the proof text follows the usual mathematical style -of concluding $P(x)$ from $\exists x. P(x)$, while carefully introducing $x$ -as a new local variable. Technically, \isakeyword{obtain} is similar to -\isakeyword{fix} and \isakeyword{assume} together with a soundness proof of -the elimination involved. - -Here is a proof of a well known tautology. -Which rule is used where? *} - -lemma assumes ex: "\x. \y. P x y" shows "\y. \x. P x y" -proof - fix y - from ex obtain x where "\y. P x y" .. - hence "P x y" .. - thus "\x. P x y" .. -qed - -subsection{*Making bigger steps*} - -text{* So far we have confined ourselves to single step proofs. Of course -powerful automatic methods can be used just as well. Here is an example, -Cantor's theorem that there is no surjective function from a set to its -powerset: *} -theorem "\S. S \ range (f :: 'a \ 'a set)" -proof - let ?S = "{x. x \ f x}" - show "?S \ range f" - proof - assume "?S \ range f" - then obtain y where fy: "?S = f y" .. - show False - proof cases - assume "y \ ?S" - with fy show False by blast - next - assume "y \ ?S" - with fy show False by blast - qed - qed -qed -text{*\noindent -For a start, the example demonstrates two new constructs: -\begin{itemize} -\item \isakeyword{let} introduces an abbreviation for a term, in our case -the witness for the claim. -\item Proof by @{text"cases"} starts a proof by cases. Note that it remains -implicit what the two cases are: it is merely expected that the two subproofs -prove @{text"P \ ?thesis"} and @{text"\P \ ?thesis"} (in that order) -for some @{term P}. -\end{itemize} -If you wonder how to \isakeyword{obtain} @{term y}: -via the predefined elimination rule @{thm rangeE[no_vars]}. - -Method @{text blast} is used because the contradiction does not follow easily -by just a single rule. If you find the proof too cryptic for human -consumption, here is a more detailed version; the beginning up to -\isakeyword{obtain} stays unchanged. *} - -theorem "\S. S \ range (f :: 'a \ 'a set)" -proof - let ?S = "{x. x \ f x}" - show "?S \ range f" - proof - assume "?S \ range f" - then obtain y where fy: "?S = f y" .. - show False - proof cases - assume "y \ ?S" - hence "y \ f y" by simp - hence "y \ ?S" by(simp add:fy) - thus False by contradiction - next - assume "y \ ?S" - hence "y \ f y" by simp - hence "y \ ?S" by(simp add:fy) - thus False by contradiction - qed - qed -qed -text{*\noindent Method @{text contradiction} succeeds if both $P$ and -$\neg P$ are among the assumptions and the facts fed into that step, in any order. - -As it happens, Cantor's theorem can be proved automatically by best-first -search. Depth-first search would diverge, but best-first search successfully -navigates through the large search space: -*} - -theorem "\S. S \ range (f :: 'a \ 'a set)" -by best -(* Of course this only works in the context of HOL's carefully -constructed introduction and elimination rules for set theory.*) - -subsection{*Raw proof blocks*} - -text{* Although we have shown how to employ powerful automatic methods like -@{text blast} to achieve bigger proof steps, there may still be the -tendency to use the default introduction and elimination rules to -decompose goals and facts. This can lead to very tedious proofs: -*} -(*<*)ML"set quick_and_dirty"(*>*) -lemma "\x y. A x y \ B x y \ C x y" -proof - fix x show "\y. A x y \ B x y \ C x y" - proof - fix y show "A x y \ B x y \ C x y" - proof - assume "A x y \ B x y" - show "C x y" sorry - qed - qed -qed -text{*\noindent Since we are only interested in the decomposition and not the -actual proof, the latter has been replaced by -\isakeyword{sorry}. Command \isakeyword{sorry} proves anything but is -only allowed in quick and dirty mode, the default interactive mode. It -is very convenient for top down proof development. - -Luckily we can avoid this step by step decomposition very easily: *} - -lemma "\x y. A x y \ B x y \ C x y" -proof - - have "\x y. \ A x y; B x y \ \ C x y" - proof - - fix x y assume "A x y" "B x y" - show "C x y" sorry - qed - thus ?thesis by blast -qed - -text{*\noindent -This can be simplified further by \emph{raw proof blocks}, i.e.\ -proofs enclosed in braces: *} - -lemma "\x y. A x y \ B x y \ C x y" -proof - - { fix x y assume "A x y" "B x y" - have "C x y" sorry } - thus ?thesis by blast -qed - -text{*\noindent The result of the raw proof block is the same theorem -as above, namely @{prop"\x y. \ A x y; B x y \ \ C x y"}. Raw -proof blocks are like ordinary proofs except that they do not prove -some explicitly stated property but that the property emerges directly -out of the \isakeyword{fixe}s, \isakeyword{assume}s and -\isakeyword{have} in the block. Thus they again serve to avoid -duplication. Note that the conclusion of a raw proof block is stated with -\isakeyword{have} rather than \isakeyword{show} because it is not the -conclusion of some pending goal but some independent claim. - -The general idea demonstrated in this subsection is very -important in Isar and distinguishes it from tactic-style proofs: -\begin{quote}\em -Do not manipulate the proof state into a particular form by applying -tactics but state the desired form explicitly and let the tactic verify -that from this form the original goal follows. -\end{quote} -This yields more readable and also more robust proofs. *} - -subsection{*Further refinements*} - -text{* This subsection discusses some further tricks that can make -life easier although they are not essential. *} - -subsubsection{*\isakeyword{and}*} - -text{* Propositions (following \isakeyword{assume} etc) may but need not be -separated by \isakeyword{and}. This is not just for readability -(\isakeyword{from} \isa{A} \isakeyword{and} \isa{B} looks nicer than -\isakeyword{from} \isa{A} \isa{B}) but for structuring lists of propositions -into possibly named blocks. In -\begin{center} -\isakeyword{assume} \isa{A:} $A_1$ $A_2$ \isakeyword{and} \isa{B:} $A_3$ -\isakeyword{and} $A_4$ -\end{center} -label \isa{A} refers to the list of propositions $A_1$ $A_2$ and -label \isa{B} to $A_3$. *} - -subsubsection{*\isakeyword{note}*} -text{* If you want to remember intermediate fact(s) that cannot be -named directly, use \isakeyword{note}. For example the result of raw -proof block can be named by following it with -\isakeyword{note}~@{text"some_name = this"}. As a side effect, -@{text this} is set to the list of facts on the right-hand side. You -can also say @{text"note some_fact"}, which simply sets @{text this}, -i.e.\ recalls @{text"some_fact"}, e.g.\ in a \isakeyword{moreover} sequence. *} - - -subsubsection{*\isakeyword{fixes}*} - -text{* Sometimes it is necessary to decorate a proposition with type -constraints, as in Cantor's theorem above. These type constraints tend -to make the theorem less readable. The situation can be improved a -little by combining the type constraint with an outer @{text"\"}: *} - -theorem "\f :: 'a \ 'a set. \S. S \ range f" -(*<*)oops(*>*) -text{*\noindent However, now @{term f} is bound and we need a -\isakeyword{fix}~\isa{f} in the proof before we can refer to @{term f}. -This is avoided by \isakeyword{fixes}: *} - -theorem fixes f :: "'a \ 'a set" shows "\S. S \ range f" -(*<*)oops(*>*) -text{* \noindent -Even better, \isakeyword{fixes} allows to introduce concrete syntax locally:*} - -lemma comm_mono: - fixes r :: "'a \ 'a \ bool" (infix ">" 60) and - f :: "'a \ 'a \ 'a" (infixl "++" 70) - assumes comm: "\x y::'a. x ++ y = y ++ x" and - mono: "\x y z::'a. x > y \ x ++ z > y ++ z" - shows "x > y \ z ++ x > z ++ y" -by(simp add: comm mono) - -text{*\noindent The concrete syntax is dropped at the end of the proof and the -theorem becomes @{thm[display,margin=60]comm_mono} -\tweakskip *} - -subsubsection{*\isakeyword{obtain}*} - -text{* The \isakeyword{obtain} construct can introduce multiple -witnesses and propositions as in the following proof fragment:*} - -lemma assumes A: "\x y. P x y \ Q x y" shows "R" -proof - - from A obtain x y where P: "P x y" and Q: "Q x y" by blast -(*<*)oops(*>*) -text{* Remember also that one does not even need to start with a formula -containing @{text"\"} as we saw in the proof of Cantor's theorem. -*} - -subsubsection{*Combining proof styles*} - -text{* Finally, whole ``scripts'' (tactic-based proofs in the style of -\cite{LNCS2283}) may appear in the leaves of the proof tree, although this is -best avoided. Here is a contrived example: *} - -lemma "A \ (A \ B) \ B" -proof - assume a: "A" - show "(A \B) \ B" - apply(rule impI) - apply(erule impE) - apply(rule a) - apply assumption - done -qed - -text{*\noindent You may need to resort to this technique if an -automatic step fails to prove the desired proposition. - -When converting a proof from tactic-style into Isar you can proceed -in a top-down manner: parts of the proof can be left in script form -while the outer structure is already expressed in Isar. *} - -(*<*)end(*>*) diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/ROOT.ML --- a/doc-src/TutorialI/IsarOverview/Isar/ROOT.ML Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2 +0,0 @@ -use_thy "Logic"; -use_thy "Induction" diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/document/.cvsignore --- a/doc-src/TutorialI/IsarOverview/Isar/document/.cvsignore Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2 +0,0 @@ -*.sty -session.tex \ No newline at end of file diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/document/intro.tex --- a/doc-src/TutorialI/IsarOverview/Isar/document/intro.tex Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,160 +0,0 @@ -\section{Introduction} - -Isabelle is a generic proof assistant. Isar is an extension of -Isabelle with structured proofs in a stylised language of mathematics. -These proofs are readable for both a human and a machine. -Isabelle/HOL~\cite{LNCS2283} is a specialisation of Isabelle with -higher-order logic (HOL). This paper is a compact introduction to -structured proofs in Isar/HOL, an extension of Isabelle/HOL. We -intentionally do not present the full language but concentrate on the -essentials. Neither do we give a formal semantics of Isar. Instead we -introduce Isar by example. We believe that the language ``speaks for -itself'' in the same way that traditional mathematical proofs do, -which are also introduced by example rather than by teaching students -logic first. A detailed exposition of Isar can be found in Markus -Wenzel's PhD thesis~\cite{Wenzel-PhD} (which also discusses related -work) and the Isar reference manual~\cite{Isar-Ref-Man}. - -\subsection{Background} - -Interactive theorem proving has been dominated by a model of proof -that goes back to the LCF system~\cite{LCF}: a proof is a more or less -structured sequence of commands that manipulate an implicit proof -state. Thus the proof text is only suitable for the machine; for a -human, the proof only comes alive when he can see the state changes -caused by the stepwise execution of the commands. Such proofs are like -uncommented assembly language programs. We call them -\emph{tactic-style} proofs because LCF proof commands were called -tactics. - -A radically different approach was taken by the Mizar -system~\cite{Rudnicki92} where proofs are written in a stylised language akin -to that used in ordinary mathematics texts. The most important argument in -favour of a mathematics-like proof language is communication: as soon as not -just the theorem but also the proof becomes an object of interest, it should -be readable. From a system development point of view there is a second -important argument against tactic-style proofs: they are much harder to -maintain when the system is modified. -%The reason is as follows. Since it is -%usually quite unclear what exactly is proved at some point in the middle of a -%command sequence, updating a failed proof often requires executing the proof -%up to the point of failure using the old version of the system. In contrast, -%mathematics-like proofs contain enough information to tell you what is proved -%at any point. - -For these reasons the Isabelle system, originally firmly in the -LCF-tradition, was extended with a language for writing structured -proofs in a mathematics-like style. As the name already indicates, -Isar was certainly inspired by Mizar. However, there are many -differences. For a start, Isar is generic: only a few of the language -constructs described below are specific to HOL; many are generic and -thus available for any logic defined in Isabelle, e.g.\ ZF. -Furthermore, we have Isabelle's powerful automatic proof procedures at -our disposal. A closer comparison of Isar and Mizar can be found -elsewhere~\cite{WenzelW-JAR}. - -\subsection{A first glimpse of Isar} -Below you find a simplified grammar for Isar proofs. -Parentheses are used for grouping and $^?$ indicates an optional item: -\begin{center} -\begin{tabular}{lrl} -\emph{proof} & ::= & \isakeyword{proof} \emph{method}$^?$ - \emph{statement}* - \isakeyword{qed} \\ - &$\mid$& \isakeyword{by} \emph{method}\\[1ex] -\emph{statement} &::= & \isakeyword{fix} \emph{variables} \\ - &$\mid$& \isakeyword{assume} \emph{propositions} \\ - &$\mid$& (\isakeyword{from} \emph{facts})$^?$ - (\isakeyword{show} $\mid$ \isakeyword{have}) - \emph{propositions} \emph{proof} \\[1ex] -\emph{proposition} &::=& (\emph{label}{\bf:})$^?$ \emph{string} \\[1ex] -\emph{fact} &::=& \emph{label} -\end{tabular} -\end{center} -A proof can be either compound (\isakeyword{proof} -- -\isakeyword{qed}) or atomic (\isakeyword{by}). A \emph{method} is a -proof method (tactic) offered by the underlying theorem prover. -Thus this grammar is generic both w.r.t.\ the logic and the theorem prover. - -This is a typical proof skeleton: -\begin{center} -\begin{tabular}{@{}ll} -\isakeyword{proof}\\ -\hspace*{3ex}\isakeyword{assume} \isa{"}\emph{the-assm}\isa{"}\\ -\hspace*{3ex}\isakeyword{have} \isa{"}\dots\isa{"} & --- intermediate result\\ -\hspace*{3ex}\vdots\\ -\hspace*{3ex}\isakeyword{have} \isa{"}\dots\isa{"} & --- intermediate result\\ -\hspace*{3ex}\isakeyword{show} \isa{"}\emph{the-concl}\isa{"}\\ -\isakeyword{qed} -\end{tabular} -\end{center} -It proves \emph{the-assm}~$\Longrightarrow$~\emph{the-concl}. Text starting with -``---'' is a comment. The intermediate \isakeyword{have}s are only -there to bridge the gap between the assumption and the conclusion and -do not contribute to the theorem being proved. In contrast, -\isakeyword{show} establishes the conclusion of the theorem. - -\subsection{Bits of Isabelle} - -We recall some basic notions and notation from Isabelle. For more -details and for instructions how to run examples see -elsewhere~\cite{LNCS2283}. - -Isabelle's meta-logic comes with a type of \emph{propositions} with -implication $\Longrightarrow$ and a universal quantifier $\bigwedge$ for expressing -inference rules and generality. Iterated implications $A_1 \Longrightarrow \dots -A_n \Longrightarrow A$ may be abbreviated to $[\![ A_1; \dots; A_n ]\!] \Longrightarrow A$. -Applying a theorem $A \Longrightarrow B$ (named \isa{T}) to a theorem $A$ (named -\isa{U}) is written \isa{T[OF U]} and yields theorem $B$. - -Isabelle terms are simply typed. Function types are -written $\tau_1 \Rightarrow \tau_2$. - -Free variables that may be instantiated (``logical variables'' in Prolog -parlance) are prefixed with a \isa{?}. Typically, theorems are stated with -ordinary free variables but after the proof those are automatically replaced -by \isa{?}-variables. Thus the theorem can be used with arbitrary instances -of its free variables. - -Isabelle/HOL offers all the usual logical symbols like $\longrightarrow$, $\land$, -$\forall$ etc. HOL formulae are propositions, e.g.\ $\forall$ can appear below -$\Longrightarrow$, but not the other way around. Beware that $\longrightarrow$ binds more -tightly than $\Longrightarrow$: in $\forall x. P \longrightarrow Q$ the $\forall x$ covers $P \longrightarrow Q$, whereas -in $\forall x. P \Longrightarrow Q$ it covers only $P$. - -Proof methods include \isa{rule} (which performs a backwards -step with a given rule, unifying the conclusion of the rule with the -current subgoal and replacing the subgoal by the premises of the -rule), \isa{simp} (for simplification) and \isa{blast} (for predicate -calculus reasoning). - -\subsection{Overview of the paper} - -The rest of the paper is divided into two parts. -Section~\ref{sec:Logic} introduces proofs in pure logic based on -natural deduction. Section~\ref{sec:Induct} is dedicated to induction, -the key reasoning principle for computer science applications. - -There are two further areas where Isar provides specific support, but -which we do not document here. Reasoning by chains of (in)equations is -described elsewhere~\cite{BauerW-TPHOLs01}. Reasoning about -axiomatically defined structures by means of so called ``locales'' was -first described in \cite{KWP-TPHOLs99} but has evolved much since -then. - -Finally, a word of warning for potential writers of Isar proofs. It -has always been easier to write obscure rather than readable texts. -Similarly, tactic-style proofs are often (though by no means always!) -easier to write than readable ones: structure does not emerge -automatically but needs to be understood and imposed. If the precise -structure of the proof is unclear at beginning, it can be useful to -start in a tactic-based style for exploratory purposes until one has -found a proof which can be converted into a structured text in a -second step. -% Top down conversion is possible because Isar -%allows tactic-style proofs as components of structured ones. - -%Do not be mislead by the simplicity of the formulae being proved, -%especially in the beginning. Isar has been used very successfully in -%large applications, for example the formalisation of Java, in -%particular the verification of the bytecode verifier~\cite{KleinN-TCS}. diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/document/llncs.cls --- a/doc-src/TutorialI/IsarOverview/Isar/document/llncs.cls Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1189 +0,0 @@ -% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002) -% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science -% -%% -%% \CharacterTable -%% {Upper-case \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z -%% Lower-case \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z -%% Digits \0\1\2\3\4\5\6\7\8\9 -%% Exclamation \! Double quote \" Hash (number) \# -%% Dollar \$ Percent \% Ampersand \& -%% Acute accent \' Left paren \( Right paren \) -%% Asterisk \* Plus \+ Comma \, -%% Minus \- Point \. Solidus \/ -%% Colon \: Semicolon \; Less than \< -%% Equals \= Greater than \> Question mark \? -%% Commercial at \@ Left bracket \[ Backslash \\ -%% Right bracket \] Circumflex \^ Underscore \_ -%% Grave accent \` Left brace \{ Vertical bar \| -%% Right brace \} Tilde \~} -%% -\NeedsTeXFormat{LaTeX2e}[1995/12/01] -\ProvidesClass{llncs}[2002/01/28 v2.13 -^^J LaTeX document class for Lecture Notes in Computer Science] -% Options -\let\if@envcntreset\iffalse -\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue} -\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y} -\DeclareOption{oribibl}{\let\oribibl=Y} -\let\if@custvec\iftrue -\DeclareOption{orivec}{\let\if@custvec\iffalse} -\let\if@envcntsame\iffalse -\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue} -\let\if@envcntsect\iffalse -\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue} -\let\if@runhead\iffalse -\DeclareOption{runningheads}{\let\if@runhead\iftrue} - -\let\if@openbib\iffalse -\DeclareOption{openbib}{\let\if@openbib\iftrue} - -% languages -\let\switcht@@therlang\relax -\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}} -\def\ds@francais{\def\switcht@@therlang{\switcht@francais}} - -\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}} - -\ProcessOptions - -\LoadClass[twoside]{article} -\RequirePackage{multicol} % needed for the list of participants, index - -\setlength{\textwidth}{12.2cm} -\setlength{\textheight}{19.3cm} -\renewcommand\@pnumwidth{2em} -\renewcommand\@tocrmarg{3.5em} -% -\def\@dottedtocline#1#2#3#4#5{% - \ifnum #1>\c@tocdepth \else - \vskip \z@ \@plus.2\p@ - {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm - \parfillskip -\rightskip \pretolerance=10000 - \parindent #2\relax\@afterindenttrue - \interlinepenalty\@M - \leavevmode - \@tempdima #3\relax - \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip - {#4}\nobreak - \leaders\hbox{$\m@th - \mkern \@dotsep mu\hbox{.}\mkern \@dotsep - mu$}\hfill - \nobreak - \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}% - \par}% - \fi} -% -\def\switcht@albion{% -\def\abstractname{Abstract.} -\def\ackname{Acknowledgement.} -\def\andname{and} -\def\lastandname{\unskip, 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}} -\switcht@albion -% Names of theorem like environments are already defined -% but must be translated if another language is chosen -% -% French section -\def\switcht@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{Preuve}% - \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\switcht@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}% -} - -% Ragged bottom for the actual page -\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil -\global\let\@textbottom\relax}} - -\renewcommand\small{% - \@setfontsize\small\@ixpt{11}% - \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@ - \abovedisplayshortskip \z@ \@plus2\p@ - \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@ - \def\@listi{\leftmargin\leftmargini - \parsep 0\p@ \@plus1\p@ \@minus\p@ - \topsep 8\p@ \@plus2\p@ \@minus4\p@ - \itemsep0\p@}% - \belowdisplayskip \abovedisplayskip -} - -\frenchspacing -\widowpenalty=10000 -\clubpenalty=10000 - -\setlength\oddsidemargin {63\p@} -\setlength\evensidemargin {63\p@} -\setlength\marginparwidth {90\p@} - -\setlength\headsep {16\p@} - -\setlength\footnotesep{7.7\p@} -\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@} -\setlength\intextsep {8mm\@plus 2\p@ \@minus 2\p@} - -\setcounter{secnumdepth}{2} - -\newcounter {chapter} -\renewcommand\thechapter {\@arabic\c@chapter} - -\newif\if@mainmatter \@mainmattertrue -\newcommand\frontmatter{\cleardoublepage - \@mainmatterfalse\pagenumbering{Roman}} -\newcommand\mainmatter{\cleardoublepage - \@mainmattertrue\pagenumbering{arabic}} -\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi - \@mainmatterfalse} - -\renewcommand\part{\cleardoublepage - \thispagestyle{empty}% - \if@twocolumn - \onecolumn - \@tempswatrue - \else - \@tempswafalse - \fi - \null\vfil - \secdef\@part\@spart} - -\def\@part[#1]#2{% - \ifnum \c@secnumdepth >-2\relax - \refstepcounter{part}% - \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}% - \else - \addcontentsline{toc}{part}{#1}% - \fi - \markboth{}{}% - {\centering - \interlinepenalty \@M - \normalfont - \ifnum \c@secnumdepth >-2\relax - \huge\bfseries \partname~\thepart - \par - \vskip 20\p@ - \fi - \Huge \bfseries #2\par}% - \@endpart} -\def\@spart#1{% - {\centering - \interlinepenalty \@M - \normalfont - \Huge \bfseries #1\par}% - \@endpart} -\def\@endpart{\vfil\newpage - \if@twoside - \null - \thispagestyle{empty}% - \newpage - \fi - \if@tempswa - \twocolumn - \fi} - -\newcommand\chapter{\clearpage - \thispagestyle{empty}% - \global\@topnum\z@ - \@afterindentfalse - \secdef\@chapter\@schapter} -\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne - \if@mainmatter - \refstepcounter{chapter}% - \typeout{\@chapapp\space\thechapter.}% - \addcontentsline{toc}{chapter}% - {\protect\numberline{\thechapter}#1}% - \else - \addcontentsline{toc}{chapter}{#1}% - \fi - \else - \addcontentsline{toc}{chapter}{#1}% - \fi - \chaptermark{#1}% - \addtocontents{lof}{\protect\addvspace{10\p@}}% - \addtocontents{lot}{\protect\addvspace{10\p@}}% - \if@twocolumn - \@topnewpage[\@makechapterhead{#2}]% - \else - \@makechapterhead{#2}% - \@afterheading - \fi} -\def\@makechapterhead#1{% -% \vspace*{50\p@}% - {\centering - \ifnum \c@secnumdepth >\m@ne - \if@mainmatter - \large\bfseries \@chapapp{} \thechapter - \par\nobreak - \vskip 20\p@ - \fi - \fi - \interlinepenalty\@M - \Large \bfseries #1\par\nobreak - \vskip 40\p@ - }} -\def\@schapter#1{\if@twocolumn - \@topnewpage[\@makeschapterhead{#1}]% - \else - \@makeschapterhead{#1}% - \@afterheading - \fi} -\def\@makeschapterhead#1{% -% \vspace*{50\p@}% - {\centering - \normalfont - \interlinepenalty\@M - \Large \bfseries #1\par\nobreak - \vskip 40\p@ - }} - -\renewcommand\section{\@startsection{section}{1}{\z@}% - {-18\p@ \@plus -4\p@ \@minus -4\p@}% - {12\p@ \@plus 4\p@ \@minus 4\p@}% - {\normalfont\large\bfseries\boldmath - \rightskip=\z@ \@plus 8em\pretolerance=10000 }} -\renewcommand\subsection{\@startsection{subsection}{2}{\z@}% - {-18\p@ \@plus -4\p@ \@minus -4\p@}% - {8\p@ \@plus 4\p@ \@minus 4\p@}% - {\normalfont\normalsize\bfseries\boldmath - \rightskip=\z@ \@plus 8em\pretolerance=10000 }} -\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}% - {-18\p@ \@plus -4\p@ \@minus -4\p@}% - {-0.5em \@plus -0.22em \@minus -0.1em}% - {\normalfont\normalsize\bfseries\boldmath}} -\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}% - {-12\p@ \@plus -4\p@ \@minus -4\p@}% - {-0.5em \@plus -0.22em \@minus -0.1em}% - {\normalfont\normalsize\itshape}} -\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use - \string\subparagraph\space with this class}\vskip0.5cm -You should not use \verb|\subparagraph| with this class.\vskip0.5cm} - -\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00} -\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01} -\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02} -\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03} -\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04} -\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05} -\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06} -\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07} -\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08} -\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09} -\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A} - -\let\footnotesize\small - -\if@custvec -\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}} -{\mbox{\boldmath$\textstyle#1$}} -{\mbox{\boldmath$\scriptstyle#1$}} -{\mbox{\boldmath$\scriptscriptstyle#1$}}} -\fi - -\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}} -\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil -\penalty50\hskip1em\null\nobreak\hfil\squareforqed -\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi} - -\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip -\halign{\hfil -$\displaystyle##$\hfil\cr\gets\cr\to\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets -\cr\to\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets -\cr\to\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr -\gets\cr\to\cr}}}}} -\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil -$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr -\noalign{\vskip1.2pt}=\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr -\noalign{\vskip1pt}=\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr -<\cr -\noalign{\vskip0.9pt}=\cr}}}}} -\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil -$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr -\noalign{\vskip1.2pt}=\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr -\noalign{\vskip1pt}=\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr ->\cr -\noalign{\vskip0.9pt}=\cr}}}}} -\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip -\halign{\hfil -$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr ->\cr\noalign{\vskip-1pt}<\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr ->\cr\noalign{\vskip-0.8pt}<\cr}}} -{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr ->\cr\noalign{\vskip-0.3pt}<\cr}}}}} -\def\bbbr{{\rm I\!R}} %reelle Zahlen -\def\bbbm{{\rm I\!M}} -\def\bbbn{{\rm I\!N}} %natuerliche Zahlen -\def\bbbf{{\rm I\!F}} -\def\bbbh{{\rm I\!H}} -\def\bbbk{{\rm I\!K}} -\def\bbbp{{\rm I\!P}} -\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l} -{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}} -\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox -to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} -{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox -to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} -{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox -to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}} -{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox -to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}} -\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm -Q$}\hbox{\raise -0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}} -{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise -0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}} -{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise -0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}} -{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise -0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}} -\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm -T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} -{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox -to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} -{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox -to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}} -{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox -to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}} -\def\bbbs{{\mathchoice -{\setbox0=\hbox{$\displaystyle \rm S$}\hbox{\raise0.5\ht0\hbox -to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox -to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}} -{\setbox0=\hbox{$\textstyle \rm S$}\hbox{\raise0.5\ht0\hbox -to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox -to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}} -{\setbox0=\hbox{$\scriptstyle \rm S$}\hbox{\raise0.5\ht0\hbox -to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox -to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}} -{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox -to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox -to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}} -\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}} -{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}} -{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}} -{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}} - -\let\ts\, - -\setlength\leftmargini {17\p@} -\setlength\leftmargin {\leftmargini} -\setlength\leftmarginii {\leftmargini} -\setlength\leftmarginiii {\leftmargini} -\setlength\leftmarginiv {\leftmargini} -\setlength \labelsep {.5em} -\setlength \labelwidth{\leftmargini} -\addtolength\labelwidth{-\labelsep} - -\def\@listI{\leftmargin\leftmargini - \parsep 0\p@ \@plus1\p@ \@minus\p@ - \topsep 8\p@ \@plus2\p@ \@minus4\p@ - \itemsep0\p@} -\let\@listi\@listI -\@listi -\def\@listii {\leftmargin\leftmarginii - \labelwidth\leftmarginii - \advance\labelwidth-\labelsep - \topsep 0\p@ \@plus2\p@ \@minus\p@} -\def\@listiii{\leftmargin\leftmarginiii - \labelwidth\leftmarginiii - \advance\labelwidth-\labelsep - \topsep 0\p@ \@plus\p@\@minus\p@ - \parsep \z@ - \partopsep \p@ \@plus\z@ \@minus\p@} - -\renewcommand\labelitemi{\normalfont\bfseries --} -\renewcommand\labelitemii{$\m@th\bullet$} - -\setlength\arraycolsep{1.4\p@} -\setlength\tabcolsep{1.4\p@} - -\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}% - {{\contentsname}}} - \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}} - \def\lastand{\ifnum\value{auco}=2\relax - \unskip{} \andname\ - \else - \unskip \lastandname\ - \fi}% - \def\and{\stepcounter{@auth}\relax - \ifnum\value{@auth}=\value{auco}% - \lastand - \else - \unskip, - \fi}% - \@starttoc{toc}\if@restonecol\twocolumn\fi} - -\def\l@part#1#2{\addpenalty{\@secpenalty}% - \addvspace{2em plus\p@}% % space above part line - \begingroup - \parindent \z@ - \rightskip \z@ plus 5em - \hrule\vskip5pt - \large % same size as for a contribution heading - \bfseries\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\hyperhrefextend{\ifx\hyper@anchor\@undefined\else -{chapter.\thechapter}\fi} - -\def\addnumcontentsmark#1#2#3{% -\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline - {\thechapter}#3}{\thepage}\hyperhrefextend}} -\def\addcontentsmark#1#2#3{% -\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}} -\def\addcontentsmarkwop#1#2#3{% -\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}} - -\def\@adcmk[#1]{\ifcase #1 \or -\def\@gtempa{\addnumcontentsmark}% - \or \def\@gtempa{\addcontentsmark}% - \or \def\@gtempa{\addcontentsmarkwop}% - \fi\@gtempa{toc}{chapter}} -\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}} - -\def\l@chapter#1#2{\addpenalty{-\@highpenalty} - \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup - \parindent \z@ \rightskip \@tocrmarg - \advance\rightskip by 0pt plus 2cm - \parfillskip -\rightskip \pretolerance=10000 - \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip - {\large\bfseries\boldmath#1}\ifx0#2\hfil\null - \else - \nobreak - \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern - \@dotsep mu$}\hfill - \nobreak\hbox to\@pnumwidth{\hss #2}% - \fi\par - \penalty\@highpenalty \endgroup} - -\def\l@title#1#2{\addpenalty{-\@highpenalty} - \addvspace{8pt plus 1pt} - \@tempdima \z@ - \begingroup - \parindent \z@ \rightskip \@tocrmarg - \advance\rightskip by 0pt plus 2cm - \parfillskip -\rightskip \pretolerance=10000 - \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip - #1\nobreak - \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern - \@dotsep mu$}\hfill - \nobreak\hbox to\@pnumwidth{\hss #2}\par - \penalty\@highpenalty \endgroup} - -\def\l@author#1#2{\addpenalty{\@highpenalty} - \@tempdima=\z@ %15\p@ - \begingroup - \parindent \z@ \rightskip \@tocrmarg - \advance\rightskip by 0pt plus 2cm - \pretolerance=10000 - \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip - \textit{#1}\par - \penalty\@highpenalty \endgroup} - -\setcounter{tocdepth}{0} -\newdimen\tocchpnum -\newdimen\tocsecnum -\newdimen\tocsectotal -\newdimen\tocsubsecnum -\newdimen\tocsubsectotal -\newdimen\tocsubsubsecnum -\newdimen\tocsubsubsectotal -\newdimen\tocparanum -\newdimen\tocparatotal -\newdimen\tocsubparanum -\tocchpnum=\z@ % no chapter numbers -\tocsecnum=15\p@ % section 88. plus 2.222pt -\tocsubsecnum=23\p@ % subsection 88.8 plus 2.222pt -\tocsubsubsecnum=27\p@ % subsubsection 88.8.8 plus 1.444pt -\tocparanum=35\p@ % paragraph 88.8.8.8 plus 1.666pt -\tocsubparanum=43\p@ % subparagraph 88.8.8.8.8 plus 1.888pt -\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\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}} - \@starttoc{lof}\if@restonecol\twocolumn\fi} -\def\l@figure{\@dottedtocline{1}{0em}{1.5em}} - -\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn - \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}} - \@starttoc{lot}\if@restonecol\twocolumn\fi} -\let\l@table\l@figure - -\renewcommand\listoffigures{% - \section*{\listfigurename - \@mkboth{\listfigurename}{\listfigurename}}% - \@starttoc{lof}% - } - -\renewcommand\listoftables{% - \section*{\listtablename - \@mkboth{\listtablename}{\listtablename}}% - \@starttoc{lot}% - } - -\ifx\oribibl\undefined -\ifx\citeauthoryear\undefined -\renewenvironment{thebibliography}[1] - {\section*{\refname} - \def\@biblabel##1{##1.} - \small - \list{\@biblabel{\@arabic\c@enumiv}}% - {\settowidth\labelwidth{\@biblabel{#1}}% - \leftmargin\labelwidth - \advance\leftmargin\labelsep - \if@openbib - \advance\leftmargin\bibindent - \itemindent -\bibindent - \listparindent \itemindent - \parsep \z@ - \fi - \usecounter{enumiv}% - \let\p@enumiv\@empty - \renewcommand\theenumiv{\@arabic\c@enumiv}}% - \if@openbib - \renewcommand\newblock{\par}% - \else - \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}% - \fi - \sloppy\clubpenalty4000\widowpenalty4000% - \sfcode`\.=\@m} - {\def\@noitemerr - {\@latex@warning{Empty `thebibliography' environment}}% - \endlist} -\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw - {\let\protect\noexpand\immediate - \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces} -\newcount\@tempcntc -\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi - \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do - {\@ifundefined - {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries - ?}\@warning - {Citation `\@citeb' on page \thepage \space undefined}}% - {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}% - \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne - \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}% - \else - \advance\@tempcntb\@ne - \ifnum\@tempcntb=\@tempcntc - \else\advance\@tempcntb\m@ne\@citeo - \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}} -\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else - \@citea\def\@citea{,\,\hskip\z@skip}% - \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else - {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else - \def\@citea{--}\fi - \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi} -\else -\renewenvironment{thebibliography}[1] - {\section*{\refname} - \small - \list{}% - {\settowidth\labelwidth{}% - \leftmargin\parindent - \itemindent=-\parindent - \labelsep=\z@ - \if@openbib - \advance\leftmargin\bibindent - \itemindent -\bibindent - \listparindent \itemindent - \parsep \z@ - \fi - \usecounter{enumiv}% - \let\p@enumiv\@empty - \renewcommand\theenumiv{}}% - \if@openbib - \renewcommand\newblock{\par}% - \else - \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}% - \fi - \sloppy\clubpenalty4000\widowpenalty4000% - \sfcode`\.=\@m} - {\def\@noitemerr - {\@latex@warning{Empty `thebibliography' environment}}% - \endlist} - \def\@cite#1{#1}% - \def\@lbibitem[#1]#2{\item[]\if@filesw - {\def\protect##1{\string ##1\space}\immediate - \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces} - \fi -\else -\@cons\@openbib@code{\noexpand\small} -\fi - -\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\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax} - -\renewenvironment{theindex} - {\@mkboth{\indexname}{\indexname}% - \thispagestyle{empty}\parindent\z@ - \parskip\z@ \@plus .3\p@\relax - \let\item\par - \def\,{\relax\ifmmode\mskip\thinmuskip - \else\hskip0.2em\ignorespaces\fi}% - \normalfont\small - \begin{multicols}{2}[\@makeschapterhead{\indexname}]% - } - {\end{multicols}} - -\renewcommand\footnoterule{% - \kern-3\p@ - \hrule\@width 2truecm - \kern2.6\p@} - \newdimen\fnindent - \fnindent1em -\long\def\@makefntext#1{% - \parindent \fnindent% - \leftskip \fnindent% - \noindent - \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1} - -\long\def\@makecaption#1#2{% - \vskip\abovecaptionskip - \sbox\@tempboxa{{\bfseries #1.} #2}% - \ifdim \wd\@tempboxa >\hsize - {\bfseries #1.} #2\par - \else - \global \@minipagefalse - \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}% - \fi - \vskip\belowcaptionskip} - -\def\fps@figure{htbp} -\def\fnum@figure{\figurename\thinspace\thefigure} -\def \@floatboxreset {% - \reset@font - \small - \@setnobreak - \@setminipage -} -\def\fps@table{htbp} -\def\fnum@table{\tablename~\thetable} -\renewenvironment{table} - {\setlength\abovecaptionskip{0\p@}% - \setlength\belowcaptionskip{10\p@}% - \@float{table}} - {\end@float} -\renewenvironment{table*} - {\setlength\abovecaptionskip{0\p@}% - \setlength\belowcaptionskip{10\p@}% - \@dblfloat{table}} - {\end@dblfloat} - -\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} - -% LaTeX does not provide a command to enter the authors institute -% addresses. The \institute command is defined here. - -\newcounter{@inst} -\newcounter{@auth} -\newcounter{auco} -\newdimen\instindent -\newbox\authrun -\newtoks\authorrunning -\newtoks\tocauthor -\newbox\titrun -\newtoks\titlerunning -\newtoks\toctitle - -\def\clearheadinfo{\gdef\@author{No Author Given}% - \gdef\@title{No Title Given}% - \gdef\@subtitle{}% - \gdef\@institute{No Institute Given}% - \gdef\@thanks{}% - \global\titlerunning={}\global\authorrunning={}% - \global\toctitle={}\global\tocauthor={}} - -\def\institute#1{\gdef\@institute{#1}} - -\def\institutename{\par - \begingroup - \parskip=\z@ - \parindent=\z@ - \setcounter{@inst}{1}% - \def\and{\par\stepcounter{@inst}% - \noindent$^{\the@inst}$\enspace\ignorespaces}% - \setbox0=\vbox{\def\thanks##1{}\@institute}% - \ifnum\c@@inst=1\relax - \gdef\fnnstart{0}% - \else - \xdef\fnnstart{\c@@inst}% - \setcounter{@inst}{1}% - \noindent$^{\the@inst}$\enspace - \fi - \ignorespaces - \@institute\par - \endgroup} - -\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or - {\star\star\star}\or \dagger\or \ddagger\or - \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger - \or \ddagger\ddagger \else\@ctrerr\fi}} - -\def\inst#1{\unskip$^{#1}$} -\def\fnmsep{\unskip$^,$} -\def\email#1{{\tt#1}} -\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}% -\@ifpackageloaded{babel}{% -\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}% -\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}% -\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}% -}{\switcht@@therlang}% -} -\def\homedir{\~{ }} - -\def\subtitle#1{\gdef\@subtitle{#1}} -\clearheadinfo - -\renewcommand\maketitle{\newpage - \refstepcounter{chapter}% - \stepcounter{section}% - \setcounter{section}{0}% - \setcounter{subsection}{0}% - \setcounter{figure}{0} - \setcounter{table}{0} - \setcounter{equation}{0} - \setcounter{footnote}{0}% - \begingroup - \parindent=\z@ - \renewcommand\thefootnote{\@fnsymbol\c@footnote}% - \if@twocolumn - \ifnum \col@number=\@ne - \@maketitle - \else - \twocolumn[\@maketitle]% - \fi - \else - \newpage - \global\@topnum\z@ % Prevents figures from going at top of page. - \@maketitle - \fi - \thispagestyle{empty}\@thanks -% - \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}% - \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}% - \instindent=\hsize - \advance\instindent by-\headlineindent - \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else - \addcontentsline{toc}{title}{\the\toctitle}\fi - \if@runhead - \if!\the\titlerunning!\else - \edef\@title{\the\titlerunning}% - \fi - \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}% - \ifdim\wd\titrun>\instindent - \typeout{Title too long for running head. Please supply}% - \typeout{a shorter form with \string\titlerunning\space prior to - \string\maketitle}% - \global\setbox\titrun=\hbox{\small\rm - Title Suppressed Due to Excessive Length}% - \fi - \xdef\@title{\copy\titrun}% - \fi -% - \if!\the\tocauthor!\relax - {\def\and{\noexpand\protect\noexpand\and}% - \protected@xdef\toc@uthor{\@author}}% - \else - \def\\{\noexpand\protect\noexpand\newline}% - \protected@xdef\scratch{\the\tocauthor}% - \protected@xdef\toc@uthor{\scratch}% - \fi - \addcontentsline{toc}{author}{\toc@uthor}% - \if@runhead - \if!\the\authorrunning! - \value{@inst}=\value{@auth}% - \setcounter{@auth}{1}% - \else - \edef\@author{\the\authorrunning}% - \fi - \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}% - \ifdim\wd\authrun>\instindent - \typeout{Names of authors too long for running head. Please supply}% - \typeout{a shorter form with \string\authorrunning\space prior to - \string\maketitle}% - \global\setbox\authrun=\hbox{\small\rm - Authors Suppressed Due to Excessive Length}% - \fi - \xdef\@author{\copy\authrun}% - \markboth{\@author}{\@title}% - \fi - \endgroup - \setcounter{footnote}{\fnnstart}% - \clearheadinfo} -% -\def\@maketitle{\newpage - \markboth{}{}% - \def\lastand{\ifnum\value{@inst}=2\relax - \unskip{} \andname\ - \else - \unskip \lastandname\ - \fi}% - \def\and{\stepcounter{@auth}\relax - \ifnum\value{@auth}=\value{@inst}% - \lastand - \else - \unskip, - \fi}% - \begin{center}% - \let\newline\\ - {\Large \bfseries\boldmath - \pretolerance=10000 - \@title \par}\vskip .8cm -\if!\@subtitle!\else {\large \bfseries\boldmath - \vskip -.65cm - \pretolerance=10000 - \@subtitle \par}\vskip .8cm\fi - \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}% - \def\thanks##1{}\@author}% - \global\value{@inst}=\value{@auth}% - \global\value{auco}=\value{@auth}% - \setcounter{@auth}{1}% -{\lineskip .5em -\noindent\ignorespaces -\@author\vskip.35cm} - {\small\institutename} - \end{center}% - } - -% 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. -% -%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% - -\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\ }} - -\if@envcntsect - \def\@thmcountersep{.} - \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape} -\else - \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 % alle Umgebungen wie Theorem. - \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}} -\else % alle Umgebungen mit eigenem Zaehler - \if@envcntsect % mit section numeriert - \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}} - \else % nicht mit section numeriert - \if@envcntreset - \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4} - \@addtoreset{#1}{section}} - \else - \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4} - \@addtoreset{#1}{chapter}}% - \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} -\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily} -\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape} -\spn@wtheorem{note}{Note}{\itshape}{\rmfamily} -\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily} -\spn@wtheorem{property}{Property}{\itshape}{\rmfamily} -\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape} -\spn@wtheorem{question}{Question}{\itshape}{\rmfamily} -\spn@wtheorem{solution}{Solution}{\itshape}{\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\ }} - } - -\renewenvironment{abstract}{% - \list{}{\advance\topsep by0.35cm\relax\small - \leftmargin=1cm - \labelwidth=\z@ - \listparindent=\z@ - \itemindent\listparindent - \rightmargin\leftmargin}\item[\hskip\labelsep - \bfseries\abstractname]} - {\endlist} - -\newdimen\headlineindent % dimension for space between -\headlineindent=1.166cm % number and text of headings. - -\def\ps@headings{\let\@mkboth\@gobbletwo - \let\@oddfoot\@empty\let\@evenfoot\@empty - \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}% - \leftmark\hfil} - \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}% - \llap{\thepage}} - \def\chaptermark##1{}% - \def\sectionmark##1{}% - \def\subsectionmark##1{}} - -\def\ps@titlepage{\let\@mkboth\@gobbletwo - \let\@oddfoot\@empty\let\@evenfoot\@empty - \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}% - \hfil} - \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}% - \llap{\thepage}} - \def\chaptermark##1{}% - \def\sectionmark##1{}% - \def\subsectionmark##1{}} - -\if@runhead\ps@headings\else -\ps@empty\fi - -\setlength\arraycolsep{1.4\p@} -\setlength\tabcolsep{1.4\p@} - -\endinput -%end of file llncs.cls diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/document/root.bib --- a/doc-src/TutorialI/IsarOverview/Isar/document/root.bib Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,52 +0,0 @@ -@string{LNCS="Lect.\ Notes in Comp.\ Sci."} -@string{Springer="Springer-Verlag"} -@string{JAR="J. Automated Reasoning"} - -@InProceedings{BauerW-TPHOLs01, -author={Gertrud Bauer and Markus Wenzel}, -title={Calculational Reasoning Revisited --- an {Isabelle/Isar} Experience}, -booktitle={Theorem Proving in Higher Order Logics, TPHOLs 2001}, -editor={R. Boulton and P. Jackson}, -year=2001,publisher=Springer,series=LNCS,volume=2152,pages="75--90"} - -@book{LCF,author="M.C.J. Gordon and Robin Milner and C.P. Wadsworth", -title="Edinburgh {LCF}: a Mechanised Logic of Computation", -publisher=Springer,series=LNCS,volume=78,year=1979} - -@InProceedings{KWP-TPHOLs99, -author={Florian Kamm{\"u}ller and Markus Wenzel and Lawrence C. Paulson}, -title={Locales: A Sectioning Concept for {Isabelle}}, -booktitle={Theorem Proving in Higher Order Logics, TPHOLs'99}, -editor={Y. Bertot and G. Dowek and A. Hirschowitz and C. Paulin and L. Thery}, -year=1999,publisher=Springer,series=LNCS,volume=1690,pages="149--165"} - -@book{LNCS2283,author={Tobias Nipkow and Lawrence Paulson and Markus Wenzel}, -title="Isabelle/HOL --- A Proof Assistant for Higher-Order Logic", -publisher=Springer,series=LNCS,volume=2283,year=2002, -note={\url{http://www.in.tum.de/~nipkow/LNCS2283/}}} - -@article{KleinN-TCS,author={Gerwin Klein and Tobias Nipkow}, -title={Verified Bytecode Verifiers}, -journal=TCS,year=2002,note={To appear. -\url{http://www.in.tum.de/~nipkow/pubs/tcs03.html}}} - -@InProceedings{Rudnicki92,author={P. Rudnicki}, -title={An Overview of the {Mizar} Project}, -booktitle={Workshop on Types for Proofs and Programs}, -year=1992,organization={Chalmers University of Technology}} - -@manual{Isar-Ref-Man,author="Markus Wenzel", -title="The Isabelle/Isar Reference Manual", -organization={Technische Universit{\"a}t M{\"u}nchen},year=2002, -note={\url{http://isabelle.in.tum.de/dist/Isabelle2002/doc/isar-ref.pdf}}} - -@phdthesis{Wenzel-PhD,author={Markus Wenzel},title={Isabelle/Isar --- A -Versatile Environment for Human-Readable Formal Proof Documents}, -school={Institut f{\"u}r Informatik, Technische Universit{\"a}t M{\"u}nchen}, -year=2002, -note={\url{http://tumb1.biblio.tu-muenchen.de/publ/diss/in/2002/wenzel.html}}} - -@article{WenzelW-JAR,author={Markus Wenzel and Freek Wiedijk}, -title={A comparison of the mathematical proof languages {Mizar} and {Isar}}, -journal=JAR,year=2003,note={To appear}} - diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/document/root.tex --- a/doc-src/TutorialI/IsarOverview/Isar/document/root.tex Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,50 +0,0 @@ -\documentclass[envcountsame]{llncs} -%\documentclass[11pt,a4paper]{article} -\usepackage{isabelle,isabellesym,pdfsetup} - -%for best-style documents ... -\urlstyle{rm} -%\isabellestyle{it} - -\newcommand{\tweakskip}{\vspace{-\medskipamount}} -\newcommand{\Tweakskip}{\tweakskip\tweakskip} - -\pagestyle{plain} - -\begin{document} - -\title{%A Compact Introduction to -Structured Proofs in Isar/HOL} -\author{Tobias Nipkow} -\institute{Institut f{\"u}r Informatik, TU M{\"u}nchen\\ - {\small\url{http://www.in.tum.de/~nipkow/}}} -\date{} -\maketitle - -\begin{abstract} - Isar is an extension of the theorem prover Isabelle with a language - for writing human-readable structured proofs. This paper is an - introduction to the basic constructs of this language. -% It is aimed at potential users of Isar -% but also discusses the design rationals -% behind the language and its constructs. -\end{abstract} - -\input{intro.tex} -\input{Logic.tex} -\Tweakskip\Tweakskip -\input{Induction.tex} - -%\Tweakskip -\small -\paragraph{Acknowledgement} -I am deeply indebted to Markus Wenzel for conceiving Isar. Clemens Ballarin, -Gertrud Bauer, Stefan Berghofer, Gerwin Klein, Norbert Schirmer, -Markus Wenzel and Freek Wiedijk commented on and improved this paper. - -\begingroup -\bibliographystyle{plain} \small\raggedright\frenchspacing -\bibliography{root} -\endgroup - -\end{document} diff -r 75a399c2781f -r 454a2ad0c381 doc-src/TutorialI/IsarOverview/Isar/makeDemo --- a/doc-src/TutorialI/IsarOverview/Isar/makeDemo Sat May 10 20:53:02 2003 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,36 +0,0 @@ -#!/usr/bin/perl -w - -sub doit { - my ($file) = @_; - - open (FILE, $file) || die $!; - undef $/; $text = ; $/ = "\n"; - close FILE || die $!; - - $_ = $text; - - s/text_raw\{\*([^*]|\*[^}])*\*\}//sg; # actual work done here - s/text\{\*([^*]|\*[^}])*\*\}//sg; # actual work done here - s/\(\*<\*\)//sg; - s/\(\*>\*\)//sg; - s/--(\ )*"([^"])*"//sg; - s/--(\ )*\{\*([^*]|\*[^}])*\*\}//sg; - - $result = $_; - - if ($text ne $result) { - print STDERR "fixing $file\n"; -# if (! -f "$file~~") { -# rename $file, "$file~~" || die $!; -# } - open (FILE, "> Demo/$file") || die $!; - print FILE $result; - close FILE || die $!; - } -} - - -foreach $file (@ARGV) { - eval { &doit($file); }; - if ($@) { print STDERR "*** doit $file: ", $@, "\n"; } -}