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