doc-src/TutorialI/Recdef/Induction.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 12 Aug 2010 15:03:34 +0200
branchisac-from-Isabelle2009-2
changeset 37913 20e3616b2d9c
parent 16417 9bc16273c2d4
permissions -rw-r--r--
prepare reactivation of isac-update-Isa09-2
nipkow@8745
     1
(*<*)
haftmann@16417
     2
theory Induction imports examples simplification begin;
nipkow@8745
     3
(*>*)
nipkow@8745
     4
nipkow@8745
     5
text{*
nipkow@8745
     6
Assuming we have defined our function such that Isabelle could prove
nipkow@8745
     7
termination and that the recursion equations (or some suitable derived
nipkow@8745
     8
equations) are simplification rules, we might like to prove something about
nipkow@8745
     9
our function. Since the function is recursive, the natural proof principle is
nipkow@8745
    10
again induction. But this time the structural form of induction that comes
nipkow@10971
    11
with datatypes is unlikely to work well --- otherwise we could have defined the
nipkow@8745
    12
function by \isacommand{primrec}. Therefore \isacommand{recdef} automatically
nipkow@9792
    13
proves a suitable induction rule $f$@{text".induct"} that follows the
nipkow@8745
    14
recursion pattern of the particular function $f$. We call this
nipkow@8745
    15
\textbf{recursion induction}. Roughly speaking, it
nipkow@8745
    16
requires you to prove for each \isacommand{recdef} equation that the property
nipkow@8745
    17
you are trying to establish holds for the left-hand side provided it holds
nipkow@8771
    18
for all recursive calls on the right-hand side. Here is a simple example
paulson@10795
    19
involving the predefined @{term"map"} functional on lists:
nipkow@8745
    20
*}
nipkow@8745
    21
nipkow@8745
    22
lemma "map f (sep(x,xs)) = sep(f x, map f xs)";
nipkow@8745
    23
nipkow@8745
    24
txt{*\noindent
paulson@10795
    25
Note that @{term"map f xs"}
nipkow@9792
    26
is the result of applying @{term"f"} to all elements of @{term"xs"}. We prove
paulson@10795
    27
this lemma by recursion induction over @{term"sep"}:
nipkow@8745
    28
*}
nipkow@8745
    29
nipkow@8745
    30
apply(induct_tac x xs rule: sep.induct);
nipkow@8745
    31
nipkow@8745
    32
txt{*\noindent
nipkow@8745
    33
The resulting proof state has three subgoals corresponding to the three
nipkow@9792
    34
clauses for @{term"sep"}:
nipkow@10362
    35
@{subgoals[display,indent=0]}
nipkow@8745
    36
The rest is pure simplification:
nipkow@8745
    37
*}
nipkow@8745
    38
nipkow@10171
    39
apply simp_all;
nipkow@10171
    40
done
nipkow@8745
    41
nipkow@8745
    42
text{*
nipkow@8745
    43
Try proving the above lemma by structural induction, and you find that you
nipkow@8745
    44
need an additional case distinction. What is worse, the names of variables
nipkow@8745
    45
are invented by Isabelle and have nothing to do with the names in the
nipkow@9792
    46
definition of @{term"sep"}.
nipkow@8745
    47
nipkow@8745
    48
In general, the format of invoking recursion induction is
nipkow@9792
    49
\begin{quote}
nipkow@11159
    50
\isacommand{apply}@{text"(induct_tac"} $x@1 \dots x@n$ @{text"rule:"} $f$@{text".induct)"}
paulson@11428
    51
\end{quote}\index{*induct_tac (method)}%
nipkow@8745
    52
where $x@1~\dots~x@n$ is a list of free variables in the subgoal and $f$ the
nipkow@8745
    53
name of a function that takes an $n$-tuple. Usually the subgoal will
nipkow@10971
    54
contain the term $f(x@1,\dots,x@n)$ but this need not be the case. The
paulson@11458
    55
induction rules do not mention $f$ at all. Here is @{thm[source]sep.induct}:
nipkow@9723
    56
\begin{isabelle}
nipkow@9689
    57
{\isasymlbrakk}~{\isasymAnd}a.~P~a~[];\isanewline
nipkow@9689
    58
~~{\isasymAnd}a~x.~P~a~[x];\isanewline
nipkow@9689
    59
~~{\isasymAnd}a~x~y~zs.~P~a~(y~\#~zs)~{\isasymLongrightarrow}~P~a~(x~\#~y~\#~zs){\isasymrbrakk}\isanewline
nipkow@9689
    60
{\isasymLongrightarrow}~P~u~v%
nipkow@9723
    61
\end{isabelle}
paulson@11458
    62
It merely says that in order to prove a property @{term"P"} of @{term"u"} and
nipkow@9792
    63
@{term"v"} you need to prove it for the three cases where @{term"v"} is the
paulson@11458
    64
empty list, the singleton list, and the list with at least two elements.
paulson@11458
    65
The final case has an induction hypothesis:  you may assume that @{term"P"}
paulson@11458
    66
holds for the tail of that list.
nipkow@8745
    67
*}
nipkow@8745
    68
nipkow@8745
    69
(*<*)
nipkow@8745
    70
end
nipkow@8745
    71
(*>*)