doc-src/TutorialI/Misc/Itrev.thy
author nipkow
Fri, 28 Jul 2000 16:02:51 +0200
changeset 9458 c613cd06d5cf
parent 8745 13b32661dde4
child 9493 494f8cd34df7
permissions -rw-r--r--
apply. -> by
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 Itrev = Main:;
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
We define a tail-recursive version of list-reversal,
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
i.e.\ one that can be compiled into a loop:
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
     8
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     9
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    10
consts itrev :: "'a list \\<Rightarrow> 'a list \\<Rightarrow> 'a list";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
"itrev []     ys = ys"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
"itrev (x#xs) ys = itrev xs (x#ys)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
The behaviour of \isa{itrev} is simple: it reverses its first argument by
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
stacking its elements onto the second argument, and returning that second
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
argument when the first one becomes empty.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
We need to show that \isa{itrev} does indeed reverse its first argument
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
provided the second one is empty:
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    21
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
lemma "itrev xs [] = rev xs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
There is no choice as to the induction variable, and we immediately simplify:
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    27
*};
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
apply(induct_tac xs, auto);
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
Unfortunately, this is not a complete success:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
~1.~\dots~itrev~list~[]~=~rev~list~{\isasymLongrightarrow}~itrev~list~[a]~=~rev~list~@~[a]%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
Just as predicted above, the overall goal, and hence the induction
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
hypothesis, is too weak to solve the induction step because of the fixed
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
\isa{[]}. The corresponding heuristic:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
\begin{quote}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
{\em 3. Generalize goals for induction by replacing constants by variables.}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
\end{quote}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
Of course one cannot do this na\"{\i}vely: \isa{itrev xs ys = rev xs} is
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
just not true---the correct generalization is
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    45
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
(*<*)oops;(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
lemma "itrev xs ys = rev xs @ ys";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
If \isa{ys} is replaced by \isa{[]}, the right-hand side simplifies to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
\isa{rev xs}, just as required.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    53
In this particular instance it was easy to guess the right generalization,
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    54
but in more complex situations a good deal of creativity is needed. This is
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
the main source of complications in inductive proofs.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    56
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    57
Although we now have two variables, only \isa{xs} is suitable for
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    58
induction, and we repeat our above proof attempt. Unfortunately, we are still
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
not there:
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
~1.~{\isasymAnd}a~list.\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
~~~~~~~itrev~list~ys~=~rev~list~@~ys~{\isasymLongrightarrow}\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
~~~~~~~itrev~list~(a~\#~ys)~=~rev~list~@~a~\#~ys%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
The induction hypothesis is still too weak, but this time it takes no
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
intuition to generalize: the problem is that \isa{ys} is fixed throughout
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
the subgoal, but the induction hypothesis needs to be applied with
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
\isa{a \# ys} instead of \isa{ys}. Hence we prove the theorem
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
for all \isa{ys} instead of a fixed one:
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    70
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    71
(*<*)oops;(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    72
lemma "\\<forall>ys. itrev xs ys = rev xs @ ys";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    73
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
This time induction on \isa{xs} followed by simplification succeeds. This
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    76
leads to another heuristic for generalization:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    77
\begin{quote}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
{\em 4. Generalize goals for induction by universally quantifying all free
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    79
variables {\em(except the induction variable itself!)}.}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    80
\end{quote}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
This prevents trivial failures like the above and does not change the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
provability of the goal. Because it is not always required, and may even
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
complicate matters in some cases, this heuristic is often not
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    84
applied blindly.
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    85
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    87
(*<*)
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    88
by(induct_tac xs, simp, simp);
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    90
(*>*)