doc-src/TutorialI/Recdef/document/Induction.tex
author wenzelm
Thu, 01 Feb 2001 20:48:58 +0100
changeset 11023 6e6c8d1ec89e
parent 10971 6852682eaf16
child 11159 07b13770c4d6
permissions -rw-r--r--
updated
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9722
a5f86aed785b *** empty log message ***
nipkow
parents: 9721
diff changeset
     1
%
a5f86aed785b *** empty log message ***
nipkow
parents: 9721
diff changeset
     2
\begin{isabellebody}%
9924
3370f6aa3200 updated;
wenzelm
parents: 9792
diff changeset
     3
\def\isabellecontext{Induction}%
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
     4
%
2665170f104a Adding generated files
nipkow
parents:
diff changeset
     5
\begin{isamarkuptext}%
2665170f104a Adding generated files
nipkow
parents:
diff changeset
     6
Assuming we have defined our function such that Isabelle could prove
2665170f104a Adding generated files
nipkow
parents:
diff changeset
     7
termination and that the recursion equations (or some suitable derived
2665170f104a Adding generated files
nipkow
parents:
diff changeset
     8
equations) are simplification rules, we might like to prove something about
2665170f104a Adding generated files
nipkow
parents:
diff changeset
     9
our function. Since the function is recursive, the natural proof principle is
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    10
again induction. But this time the structural form of induction that comes
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10950
diff changeset
    11
with datatypes is unlikely to work well --- otherwise we could have defined the
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    12
function by \isacommand{primrec}. Therefore \isacommand{recdef} automatically
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    13
proves a suitable induction rule $f$\isa{{\isachardot}induct} that follows the
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    14
recursion pattern of the particular function $f$. We call this
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    15
\textbf{recursion induction}. Roughly speaking, it
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    16
requires you to prove for each \isacommand{recdef} equation that the property
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    17
you are trying to establish holds for the left-hand side provided it holds
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10362
diff changeset
    18
for all recursive calls on the right-hand side. Here is a simple example
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10362
diff changeset
    19
involving the predefined \isa{map} functional on lists:%
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    20
\end{isamarkuptext}%
9674
f789d2490669 updated;
wenzelm
parents: 9541
diff changeset
    21
\isacommand{lemma}\ {\isachardoublequote}map\ f\ {\isacharparenleft}sep{\isacharparenleft}x{\isacharcomma}xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep{\isacharparenleft}f\ x{\isacharcomma}\ map\ f\ xs{\isacharparenright}{\isachardoublequote}%
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    22
\begin{isamarkuptxt}%
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    23
\noindent
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10362
diff changeset
    24
Note that \isa{map\ f\ xs}
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    25
is the result of applying \isa{f} to all elements of \isa{xs}. We prove
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10362
diff changeset
    26
this lemma by recursion induction over \isa{sep}:%
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    27
\end{isamarkuptxt}%
9674
f789d2490669 updated;
wenzelm
parents: 9541
diff changeset
    28
\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ x\ xs\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}%
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    29
\begin{isamarkuptxt}%
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    30
\noindent
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    31
The resulting proof state has three subgoals corresponding to the three
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    32
clauses for \isa{sep}:
10362
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 10171
diff changeset
    33
\begin{isabelle}%
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 10171
diff changeset
    34
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}a{\isachardot}\ map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\isanewline
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 10171
diff changeset
    35
\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ x{\isachardot}\ map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\isanewline
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 10171
diff changeset
    36
\ {\isadigit{3}}{\isachardot}\ {\isasymAnd}a\ x\ y\ zs{\isachardot}\isanewline
10950
aa788fcb75a5 updated;
wenzelm
parents: 10795
diff changeset
    37
\isaindent{\ {\isadigit{3}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharparenleft}y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\isanewline
aa788fcb75a5 updated;
wenzelm
parents: 10795
diff changeset
    38
\isaindent{\ {\isadigit{3}}{\isachardot}\ \ \ \ }map\ f\ {\isacharparenleft}sep\ {\isacharparenleft}a{\isacharcomma}\ x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep\ {\isacharparenleft}f\ a{\isacharcomma}\ map\ f\ {\isacharparenleft}x\ {\isacharhash}\ y\ {\isacharhash}\ zs{\isacharparenright}{\isacharparenright}%
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9722
diff changeset
    39
\end{isabelle}
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    40
The rest is pure simplification:%
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    41
\end{isamarkuptxt}%
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9924
diff changeset
    42
\isacommand{apply}\ simp{\isacharunderscore}all\isanewline
59d6633835fa *** empty log message ***
nipkow
parents: 9924
diff changeset
    43
\isacommand{done}%
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    44
\begin{isamarkuptext}%
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    45
Try proving the above lemma by structural induction, and you find that you
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    46
need an additional case distinction. What is worse, the names of variables
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    47
are invented by Isabelle and have nothing to do with the names in the
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    48
definition of \isa{sep}.
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    49
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    50
In general, the format of invoking recursion induction is
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    51
\begin{quote}
11023
6e6c8d1ec89e updated
wenzelm
parents: 10971
diff changeset
    52
\isacommand{apply}\isa{{\isacharparenleft}induct{\isacharunderscore}tac}$x@1 \dots x@n$ \isa{rule{\isacharcolon}} $f$\isa{{\isachardot}induct{\isacharparenright}}
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    53
\end{quote}\index{*induct_tac}%
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    54
where $x@1~\dots~x@n$ is a list of free variables in the subgoal and $f$ the
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    55
name of a function that takes an $n$-tuple. Usually the subgoal will
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10950
diff changeset
    56
contain the term $f(x@1,\dots,x@n)$ but this need not be the case. The
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    57
induction rules do not mention $f$ at all. For example \isa{sep{\isachardot}induct}
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9722
diff changeset
    58
\begin{isabelle}
9689
751fde5307e4 *** empty log message ***
nipkow
parents: 9674
diff changeset
    59
{\isasymlbrakk}~{\isasymAnd}a.~P~a~[];\isanewline
751fde5307e4 *** empty log message ***
nipkow
parents: 9674
diff changeset
    60
~~{\isasymAnd}a~x.~P~a~[x];\isanewline
751fde5307e4 *** empty log message ***
nipkow
parents: 9674
diff changeset
    61
~~{\isasymAnd}a~x~y~zs.~P~a~(y~\#~zs)~{\isasymLongrightarrow}~P~a~(x~\#~y~\#~zs){\isasymrbrakk}\isanewline
751fde5307e4 *** empty log message ***
nipkow
parents: 9674
diff changeset
    62
{\isasymLongrightarrow}~P~u~v%
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9722
diff changeset
    63
\end{isabelle}
9689
751fde5307e4 *** empty log message ***
nipkow
parents: 9674
diff changeset
    64
merely says that in order to prove a property \isa{P} of \isa{u} and
751fde5307e4 *** empty log message ***
nipkow
parents: 9674
diff changeset
    65
\isa{v} you need to prove it for the three cases where \isa{v} is the
8749
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    66
empty list, the singleton list, and the list with at least two elements
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    67
(in which case you may assume it holds for the tail of that list).%
2665170f104a Adding generated files
nipkow
parents:
diff changeset
    68
\end{isamarkuptext}%
9722
a5f86aed785b *** empty log message ***
nipkow
parents: 9721
diff changeset
    69
\end{isabellebody}%
9145
9f7b8de5bfaf updated;
wenzelm
parents: 8771
diff changeset
    70
%%% Local Variables:
9f7b8de5bfaf updated;
wenzelm
parents: 8771
diff changeset
    71
%%% mode: latex
9f7b8de5bfaf updated;
wenzelm
parents: 8771
diff changeset
    72
%%% TeX-master: "root"
9f7b8de5bfaf updated;
wenzelm
parents: 8771
diff changeset
    73
%%% End: