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