doc-src/TutorialI/Recdef/Induction.thy
author paulson
Mon, 19 Mar 2001 10:37:47 +0100
changeset 11212 d06fb91f22da
parent 11159 07b13770c4d6
child 11428 332347b9b942
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
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10795
diff changeset
    11
with datatypes is unlikely to work well --- otherwise we could have defined the
8745
13b32661dde4 I wonder which files i forgot.
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$@{text".induct"} that follows the
8745
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
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10362
diff changeset
    19
involving the predefined @{term"map"} functional on lists:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
lemma "map f (sep(x,xs)) = sep(f x, map f xs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
txt{*\noindent
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10362
diff changeset
    25
Note that @{term"map f xs"}
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    26
is the result of applying @{term"f"} to all elements of @{term"xs"}. We prove
10795
9e888d60d3e5 minor edits to Chapters 1-3
paulson
parents: 10362
diff changeset
    27
this lemma by recursion induction over @{term"sep"}:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
apply(induct_tac x xs rule: sep.induct);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    31
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
The resulting proof state has three subgoals corresponding to the three
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    34
clauses for @{term"sep"}:
10362
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 10171
diff changeset
    35
@{subgoals[display,indent=0]}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
The rest is pure simplification:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
10171
59d6633835fa *** empty log message ***
nipkow
parents: 9792
diff changeset
    39
apply simp_all;
59d6633835fa *** empty log message ***
nipkow
parents: 9792
diff changeset
    40
done
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
Try proving the above lemma by structural induction, and you find that you
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
need an additional case distinction. What is worse, the names of variables
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
are invented by Isabelle and have nothing to do with the names in the
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    46
definition of @{term"sep"}.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
In general, the format of invoking recursion induction is
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    49
\begin{quote}
11159
07b13770c4d6 *** empty log message ***
nipkow
parents: 10971
diff changeset
    50
\isacommand{apply}@{text"(induct_tac"} $x@1 \dots x@n$ @{text"rule:"} $f$@{text".induct)"}
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    51
\end{quote}\index{*induct_tac}%
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
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
    53
name of a function that takes an $n$-tuple. Usually the subgoal will
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10795
diff changeset
    54
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
    55
induction rules do not mention $f$ at all. For example @{thm[source]sep.induct}
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9689
diff changeset
    56
\begin{isabelle}
9689
751fde5307e4 *** empty log message ***
nipkow
parents: 9458
diff changeset
    57
{\isasymlbrakk}~{\isasymAnd}a.~P~a~[];\isanewline
751fde5307e4 *** empty log message ***
nipkow
parents: 9458
diff changeset
    58
~~{\isasymAnd}a~x.~P~a~[x];\isanewline
751fde5307e4 *** empty log message ***
nipkow
parents: 9458
diff changeset
    59
~~{\isasymAnd}a~x~y~zs.~P~a~(y~\#~zs)~{\isasymLongrightarrow}~P~a~(x~\#~y~\#~zs){\isasymrbrakk}\isanewline
751fde5307e4 *** empty log message ***
nipkow
parents: 9458
diff changeset
    60
{\isasymLongrightarrow}~P~u~v%
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9689
diff changeset
    61
\end{isabelle}
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    62
merely says that in order to prove a property @{term"P"} of @{term"u"} and
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9723
diff changeset
    63
@{term"v"} you need to prove it for the three cases where @{term"v"} is the
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
empty list, the singleton list, and the list with at least two elements
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
(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
    66
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    70
(*>*)