*** empty log message ***
authornipkow
Wed, 30 Aug 2000 14:38:48 +0200
changeset 9742 98d3ca2c18f7
parent 9741 0502f06c2d29
child 9743 d18d5c4a1f80
*** empty log message ***
doc-src/TutorialI/Misc/case_exprs.thy
doc-src/TutorialI/Misc/document/case_exprs.tex
doc-src/TutorialI/fp.tex
doc-src/TutorialI/tutorial.tex
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/TutorialI/Misc/case_exprs.thy	Wed Aug 30 14:38:48 2000 +0200
@@ -0,0 +1,84 @@
+(*<*)
+theory case_exprs = Main:
+(*>*)
+
+subsection{*Case expressions*}
+
+text{*\label{sec:case-expressions}
+HOL also features \isaindexbold{case}-expressions for analyzing
+elements of a datatype. For example,
+\begin{quote}
+@{term[display]"case xs of [] => 1 | y#ys => y"}
+\end{quote}
+evaluates to @{term"1"} if @{term"xs"} is @{term"[]"} and to @{term"y"} if 
+@{term"xs"} is @{term"y#ys"}. (Since the result in both branches must be of
+the same type, it follows that @{term"y"} is of type @{typ"nat"} and hence
+that @{term"xs"} is of type @{typ"nat list"}.)
+
+In general, if $e$ is a term of the datatype $t$ defined in
+\S\ref{sec:general-datatype} above, the corresponding
+\isa{case}-expression analyzing $e$ is
+\[
+\begin{array}{rrcl}
+\isa{case}~e~\isa{of} & C@1~x@ {11}~\dots~x@ {1k@1} & \To & e@1 \\
+                           \vdots \\
+                           \mid & C@m~x@ {m1}~\dots~x@ {mk@m} & \To & e@m
+\end{array}
+\]
+
+\begin{warn}
+\emph{All} constructors must be present, their order is fixed, and nested
+patterns are not supported.  Violating these restrictions results in strange
+error messages.
+\end{warn}
+\noindent
+Nested patterns can be simulated by nested \isa{case}-expressions: instead
+of
+% case xs of [] => 1 | [x] => x | x#(y#zs) => y
+\begin{isabelle}
+~~~case~xs~of~[]~{\isasymRightarrow}~1~|~[x]~{\isasymRightarrow}~x~|~x~\#~y~\#~zs~{\isasymRightarrow}~y
+\end{isabelle}
+write
+\begin{quote}
+@{term[display,eta_contract=false,margin=50]"case xs of [] => 1 | x#ys => (case ys of [] => x | y#zs => y)"}
+\end{quote}
+
+Note that \isa{case}-expressions may need to be enclosed in parentheses to
+indicate their scope
+*}
+
+subsection{*Structural induction and case distinction*}
+
+text{*
+\indexbold{structural induction}
+\indexbold{induction!structural}
+\indexbold{case distinction}
+Almost all the basic laws about a datatype are applied automatically during
+simplification. Only induction is invoked by hand via \isaindex{induct_tac},
+which works for any datatype. In some cases, induction is overkill and a case
+distinction over all constructors of the datatype suffices. This is performed
+by \isaindexbold{case_tac}. A trivial example:
+*}
+
+lemma "(case xs of [] \<Rightarrow> [] | y#ys \<Rightarrow> xs) = xs";
+apply(case_tac xs);
+
+txt{*\noindent
+results in the proof state
+\begin{isabelle}
+~1.~xs~=~[]~{\isasymLongrightarrow}~(case~xs~of~[]~{\isasymRightarrow}~[]~|~y~\#~ys~{\isasymRightarrow}~xs)~=~xs\isanewline
+~2.~{\isasymAnd}a~list.~xs=a\#list~{\isasymLongrightarrow}~(case~xs~of~[]~{\isasymRightarrow}~[]~|~y\#ys~{\isasymRightarrow}~xs)~=~xs%
+\end{isabelle}
+which is solved automatically:
+*}
+
+by(auto)
+
+text{*
+Note that we do not need to give a lemma a name if we do not intend to refer
+to it explicitly in the future.
+*}
+
+(*<*)
+end
+(*>*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/TutorialI/Misc/document/case_exprs.tex	Wed Aug 30 14:38:48 2000 +0200
@@ -0,0 +1,91 @@
+%
+\begin{isabellebody}%
+%
+\isamarkupsubsection{Case expressions}
+%
+\begin{isamarkuptext}%
+\label{sec:case-expressions}
+HOL also features \isaindexbold{case}-expressions for analyzing
+elements of a datatype. For example,
+\begin{quote}
+
+\begin{isabelle}%
+case\ \mbox{xs}\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ \isadigit{1}\ {\isacharbar}\ \mbox{y}\ {\isacharhash}\ \mbox{ys}\ {\isasymRightarrow}\ \mbox{y}
+\end{isabelle}%
+
+\end{quote}
+evaluates to \isa{\isadigit{1}} if \isa{\mbox{xs}} is \isa{{\isacharbrackleft}{\isacharbrackright}} and to \isa{\mbox{y}} if 
+\isa{\mbox{xs}} is \isa{\mbox{y}\ {\isacharhash}\ \mbox{ys}}. (Since the result in both branches must be of
+the same type, it follows that \isa{\mbox{y}} is of type \isa{nat} and hence
+that \isa{\mbox{xs}} is of type \isa{nat\ list}.)
+
+In general, if $e$ is a term of the datatype $t$ defined in
+\S\ref{sec:general-datatype} above, the corresponding
+\isa{case}-expression analyzing $e$ is
+\[
+\begin{array}{rrcl}
+\isa{case}~e~\isa{of} & C@1~x@ {11}~\dots~x@ {1k@1} & \To & e@1 \\
+                           \vdots \\
+                           \mid & C@m~x@ {m1}~\dots~x@ {mk@m} & \To & e@m
+\end{array}
+\]
+
+\begin{warn}
+\emph{All} constructors must be present, their order is fixed, and nested
+patterns are not supported.  Violating these restrictions results in strange
+error messages.
+\end{warn}
+\noindent
+Nested patterns can be simulated by nested \isa{case}-expressions: instead
+of
+% case xs of [] => 1 | [x] => x | x#(y#zs) => y
+\begin{isabelle}
+~~~case~xs~of~[]~{\isasymRightarrow}~1~|~[x]~{\isasymRightarrow}~x~|~x~\#~y~\#~zs~{\isasymRightarrow}~y
+\end{isabelle}
+write
+\begin{quote}
+
+\begin{isabelle}%
+case\ \mbox{xs}\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ \isadigit{1}\isanewline
+{\isacharbar}\ \mbox{x}\ {\isacharhash}\ \mbox{ys}\ {\isasymRightarrow}\ case\ \mbox{ys}\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ \mbox{x}\ {\isacharbar}\ \mbox{y}\ {\isacharhash}\ \mbox{zs}\ {\isasymRightarrow}\ \mbox{y}
+\end{isabelle}%
+
+\end{quote}
+
+Note that \isa{case}-expressions may need to be enclosed in parentheses to
+indicate their scope%
+\end{isamarkuptext}%
+%
+\isamarkupsubsection{Structural induction and case distinction}
+%
+\begin{isamarkuptext}%
+\indexbold{structural induction}
+\indexbold{induction!structural}
+\indexbold{case distinction}
+Almost all the basic laws about a datatype are applied automatically during
+simplification. Only induction is invoked by hand via \isaindex{induct_tac},
+which works for any datatype. In some cases, induction is overkill and a case
+distinction over all constructors of the datatype suffices. This is performed
+by \isaindexbold{case_tac}. A trivial example:%
+\end{isamarkuptext}%
+\isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbar}\ y{\isacharhash}ys\ {\isasymRightarrow}\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isachardoublequote}\isanewline
+\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ xs{\isacharparenright}%
+\begin{isamarkuptxt}%
+\noindent
+results in the proof state
+\begin{isabelle}
+~1.~xs~=~[]~{\isasymLongrightarrow}~(case~xs~of~[]~{\isasymRightarrow}~[]~|~y~\#~ys~{\isasymRightarrow}~xs)~=~xs\isanewline
+~2.~{\isasymAnd}a~list.~xs=a\#list~{\isasymLongrightarrow}~(case~xs~of~[]~{\isasymRightarrow}~[]~|~y\#ys~{\isasymRightarrow}~xs)~=~xs%
+\end{isabelle}
+which is solved automatically:%
+\end{isamarkuptxt}%
+\isacommand{by}{\isacharparenleft}auto{\isacharparenright}%
+\begin{isamarkuptext}%
+Note that we do not need to give a lemma a name if we do not intend to refer
+to it explicitly in the future.%
+\end{isamarkuptext}%
+\end{isabellebody}%
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "root"
+%%% End:
--- a/doc-src/TutorialI/fp.tex	Wed Aug 30 14:02:12 2000 +0200
+++ b/doc-src/TutorialI/fp.tex	Wed Aug 30 14:38:48 2000 +0200
@@ -558,29 +558,11 @@
 \input{Misc/document/arith4.tex}%
 is not proved by simplification and requires \isa{arith}.
 
-\subsubsection{Permutative rewrite rules}
-
-A rewrite rule is \textbf{permutative} if the left-hand side and right-hand
-side are the same up to renaming of variables.  The most common permutative
-rule is commutativity: $x+y = y+x$.  Another example is $(x-y)-z = (x-z)-y$.
-Such rules are problematic because once they apply, they can be used forever.
-The simplifier is aware of this danger and treats permutative rules
-separately. For details see~\cite{isabelle-ref}.
-
 \subsubsection{Tracing}
 \indexbold{tracing the simplifier}
 
 {\makeatother\input{Misc/document/trace_simp.tex}}
 
-\subsection{How it works}
-\label{sec:SimpHow}
-
-\subsubsection{Higher-order patterns}
-
-\subsubsection{Local assumptions}
-
-\subsubsection{The preprocessor}
-
 \index{simplification|)}
 
 \section{Induction heuristics}
@@ -800,9 +782,3 @@
 \input{Recdef/document/Nested2.tex}
 
 \index{*recdef|)}
-
-\section{Advanced induction techniques}
-\label{sec:advanced-ind}
-\input{Misc/document/AdvancedInd.tex}
-
-%\input{Datatype/document/Nested2.tex}
--- a/doc-src/TutorialI/tutorial.tex	Wed Aug 30 14:02:12 2000 +0200
+++ b/doc-src/TutorialI/tutorial.tex	Wed Aug 30 14:38:48 2000 +0200
@@ -65,6 +65,7 @@
 
 \input{basics}
 \input{fp}
+\input{tricks}
 \input{appendix}
 
 \bibliographystyle{plain}