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