doc-src/TutorialI/Recdef/Induction.thy
author nipkow
Tue, 25 Apr 2000 08:09:10 +0200
changeset 8771 026f37a86ea7
parent 8745 13b32661dde4
child 9458 c613cd06d5cf
permissions -rw-r--r--
*** empty log message ***
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory Induction = examples + simplification:;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     6
Assuming we have defined our function such that Isabelle could prove
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
termination and that the recursion equations (or some suitable derived
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     8
equations) are simplification rules, we might like to prove something about
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     9
our function. Since the function is recursive, the natural proof principle is
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    10
again induction. But this time the structural form of induction that comes
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
with datatypes is unlikely to work well---otherwise we could have defined the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
function by \isacommand{primrec}. Therefore \isacommand{recdef} automatically
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
proves a suitable induction rule $f$\isa{.induct} that follows the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
recursion pattern of the particular function $f$. We call this
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
\textbf{recursion induction}. Roughly speaking, it
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
requires you to prove for each \isacommand{recdef} equation that the property
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
you are trying to establish holds for the left-hand side provided it holds
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    18
for all recursive calls on the right-hand side. Here is a simple example
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
lemma "map f (sep(x,xs)) = sep(f x, map f xs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
involving the predefined \isa{map} functional on lists: \isa{map f xs}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
is the result of applying \isa{f} to all elements of \isa{xs}. We prove
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
this lemma by recursion induction w.r.t. \isa{sep}:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
apply(induct_tac x xs rule: sep.induct);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    31
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
The resulting proof state has three subgoals corresponding to the three
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
clauses for \isa{sep}:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
~1.~{\isasymAnd}a.~map~f~(sep~(a,~[]))~=~sep~(f~a,~map~f~[])\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
~2.~{\isasymAnd}a~x.~map~f~(sep~(a,~[x]))~=~sep~(f~a,~map~f~[x])\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
~3.~{\isasymAnd}a~x~y~zs.\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
~~~~~~~map~f~(sep~(a,~y~\#~zs))~=~sep~(f~a,~map~f~(y~\#~zs))~{\isasymLongrightarrow}\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
~~~~~~~map~f~(sep~(a,~x~\#~y~\#~zs))~=~sep~(f~a,~map~f~(x~\#~y~\#~zs))%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
The rest is pure simplification:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
apply auto.;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
Try proving the above lemma by structural induction, and you find that you
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
need an additional case distinction. What is worse, the names of variables
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
are invented by Isabelle and have nothing to do with the names in the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
definition of \isa{sep}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
In general, the format of invoking recursion induction is
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    53
\begin{ttbox}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    54
apply(induct_tac \(x@1 \dots x@n\) rule: \(f\).induct)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
\end{ttbox}\index{*induct_tac}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    56
where $x@1~\dots~x@n$ is a list of free variables in the subgoal and $f$ the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    57
name of a function that takes an $n$-tuple. Usually the subgoal will
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    58
contain the term $f~x@1~\dots~x@n$ but this need not be the case. The
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
induction rules do not mention $f$ at all. For example \isa{sep.induct}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    60
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    61
{\isasymlbrakk}~{\isasymAnd}a.~?P~a~[];\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
~~{\isasymAnd}a~x.~?P~a~[x];\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
~~{\isasymAnd}a~x~y~zs.~?P~a~(y~\#~zs)~{\isasymLongrightarrow}~?P~a~(x~\#~y~\#~zs){\isasymrbrakk}\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
{\isasymLongrightarrow}~?P~?u~?v%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
merely says that in order to prove a property \isa{?P} of \isa{?u} and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
\isa{?v} you need to prove it for the three cases where \isa{?v} is the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
empty list, the singleton list, and the list with at least two elements
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
(in which case you may assume it holds for the tail of that list).
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    70
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    71
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    72
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    73
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
(*>*)