doc-src/TutorialI/Misc/Itrev.thy
author nipkow
Tue, 29 Aug 2000 16:05:13 +0200
changeset 9723 a977245dfc8a
parent 9689 751fde5307e4
child 9754 a123a64cadeb
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 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
9493
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
     5
text{* Function \isa{rev} has quadratic worst-case running time
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
     6
because it calls function \isa{\at} for each element of the list and
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
     7
\isa{\at} is linear in its first argument.  A linear time version of
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
     8
\isa{rev} reqires an extra argument where the result is accumulated
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
     9
gradually, using only \isa{\#}:*}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    10
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
consts itrev :: "'a list \\<Rightarrow> 'a list \\<Rightarrow> 'a list";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
"itrev []     ys = ys"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
"itrev (x#xs) ys = itrev xs (x#ys)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
9493
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    16
text{*\noindent The behaviour of \isa{itrev} is simple: it reverses
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    17
its first argument by stacking its elements onto the second argument,
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    18
and returning that second argument when the first one becomes
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    19
empty. Note that \isa{itrev} is tail-recursive, i.e.\ it can be
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    20
compiled into a loop.
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    21
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    22
Naturally, we would like to show that \isa{itrev} does indeed reverse
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    23
its first argument provided the second one is empty: *};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
lemma "itrev xs [] = rev xs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
There is no choice as to the induction variable, and we immediately simplify:
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    29
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
9689
751fde5307e4 *** empty log message ***
nipkow
parents: 9644
diff changeset
    31
apply(induct_tac xs, simp_all);
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    34
Unfortunately, this is not a complete success:
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9689
diff changeset
    35
\begin{isabelle}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
~1.~\dots~itrev~list~[]~=~rev~list~{\isasymLongrightarrow}~itrev~list~[a]~=~rev~list~@~[a]%
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9689
diff changeset
    37
\end{isabelle}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
Just as predicted above, the overall goal, and hence the induction
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
hypothesis, is too weak to solve the induction step because of the fixed
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
\isa{[]}. The corresponding heuristic:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
\begin{quote}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
{\em 3. Generalize goals for induction by replacing constants by variables.}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
\end{quote}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
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
    46
just not true---the correct generalization is
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    47
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
(*<*)oops;(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
lemma "itrev xs ys = rev xs @ ys";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
If \isa{ys} is replaced by \isa{[]}, the right-hand side simplifies to
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9493
diff changeset
    53
@{term"rev xs"}, just as required.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    54
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
In this particular instance it was easy to guess the right generalization,
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    56
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
    57
the main source of complications in inductive proofs.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    58
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
Although we now have two variables, only \isa{xs} is suitable for
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    60
induction, and we repeat our above proof attempt. Unfortunately, we are still
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    61
not there:
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9689
diff changeset
    62
\begin{isabelle}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
~1.~{\isasymAnd}a~list.\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
~~~~~~~itrev~list~ys~=~rev~list~@~ys~{\isasymLongrightarrow}\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
~~~~~~~itrev~list~(a~\#~ys)~=~rev~list~@~a~\#~ys%
9723
a977245dfc8a *** empty log message ***
nipkow
parents: 9689
diff changeset
    66
\end{isabelle}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
The induction hypothesis is still too weak, but this time it takes no
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
intuition to generalize: the problem is that \isa{ys} is fixed throughout
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
the subgoal, but the induction hypothesis needs to be applied with
9541
d17c0b34d5c8 *** empty log message ***
nipkow
parents: 9493
diff changeset
    70
@{term"a # ys"} instead of \isa{ys}. Hence we prove the theorem
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    71
for all \isa{ys} instead of a fixed one:
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    72
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    73
(*<*)oops;(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
lemma "\\<forall>ys. itrev xs ys = rev xs @ ys";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    76
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    77
This time induction on \isa{xs} followed by simplification succeeds. This
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
leads to another heuristic for generalization:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    79
\begin{quote}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    80
{\em 4. Generalize goals for induction by universally quantifying all free
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
variables {\em(except the induction variable itself!)}.}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
\end{quote}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
This prevents trivial failures like the above and does not change the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    84
provability of the goal. Because it is not always required, and may even
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    85
complicate matters in some cases, this heuristic is often not
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
applied blindly.
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
    87
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
    88
In general, if you have tried the above heuristics and still find your
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
    89
induction does not go through, and no obvious lemma suggests itself, you may
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
    90
need to generalize your proposition even further. This requires insight into
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
    91
the problem at hand and is beyond simple rules of thumb. In a nutshell: you
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
    92
will need to be creative. Additionally, you can read \S\ref{sec:advanced-ind}
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
    93
to learn about some advanced techniques for inductive proofs.
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    94
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    95
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    96
(*<*)
9689
751fde5307e4 *** empty log message ***
nipkow
parents: 9644
diff changeset
    97
by(induct_tac xs, simp_all);
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    98
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    99
(*>*)