doc-src/TutorialI/Misc/Itrev.thy
author chaieb
Wed, 19 May 2004 11:23:59 +0200
changeset 14758 af3b71a46a1c
parent 13081 ab4a3aef3591
child 15905 0a4cc9b113c7
permissions -rw-r--r--
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller. the tactic has also changed and allows the abstaction over fuction occurences whose type is nat or int.
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
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10795
diff changeset
     5
section{*Induction Heuristics*}
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
     6
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
     7
text{*\label{sec:InductionHeuristics}
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
     8
\index{induction heuristics|(}%
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
     9
The purpose of this section is to illustrate some simple heuristics for
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    10
inductive proofs. The first one we have already mentioned in our initial
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    11
example:
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    12
\begin{quote}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    13
\emph{Theorems about recursive functions are proved by induction.}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    14
\end{quote}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    15
In case the function has more than one argument
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    16
\begin{quote}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    17
\emph{Do induction on argument number $i$ if the function is defined by
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    18
recursion in argument number $i$.}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    19
\end{quote}
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    20
When we look at the proof of @{text"(xs@ys) @ zs = xs @ (ys@zs)"}
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    21
in \S\ref{sec:intro-proof} we find
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    22
\begin{itemize}
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    23
\item @{text"@"} is recursive in
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    24
the first argument
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    25
\item @{term xs}  occurs only as the first argument of
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    26
@{text"@"}
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    27
\item both @{term ys} and @{term zs} occur at least once as
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    28
the second argument of @{text"@"}
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    29
\end{itemize}
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    30
Hence it is natural to perform induction on~@{term xs}.
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    31
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    32
The key heuristic, and the main point of this section, is to
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    33
\emph{generalize the goal before induction}.
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    34
The reason is simple: if the goal is
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    35
too specific, the induction hypothesis is too weak to allow the induction
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
    36
step to go through. Let us illustrate the idea with an example.
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    37
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    38
Function \cdx{rev} has quadratic worst-case running time
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9754
diff changeset
    39
because it calls function @{text"@"} for each element of the list and
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9754
diff changeset
    40
@{text"@"} is linear in its first argument.  A linear time version of
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    41
@{term"rev"} reqires an extra argument where the result is accumulated
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    42
gradually, using only~@{text"#"}:
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    43
*}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
10362
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 9844
diff changeset
    45
consts itrev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list";
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
"itrev []     ys = ys"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
"itrev (x#xs) ys = itrev xs (x#ys)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    50
text{*\noindent
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    51
The behaviour of \cdx{itrev} is simple: it reverses
9493
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    52
its first argument by stacking its elements onto the second argument,
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    53
and returning that second argument when the first one becomes
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    54
empty. Note that @{term"itrev"} is tail-recursive: it can be
9493
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    55
compiled into a loop.
494f8cd34df7 *** empty log message ***
nipkow
parents: 9458
diff changeset
    56
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    57
Naturally, we would like to show that @{term"itrev"} does indeed reverse
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    58
its first argument provided the second one is empty:
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    59
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    60
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    61
lemma "itrev xs [] = rev xs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
There is no choice as to the induction variable, and we immediately simplify:
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    65
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
9689
751fde5307e4 *** empty log message ***
nipkow
parents: 9644
diff changeset
    67
apply(induct_tac xs, simp_all);
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
txt{*\noindent
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    70
Unfortunately, this attempt does not prove
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    71
the induction step:
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
    72
@{subgoals[display,indent=0,margin=70]}
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    73
The induction hypothesis is too weak.  The fixed
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    74
argument,~@{term"[]"}, prevents it from rewriting the conclusion.  
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    75
This example suggests a heuristic:
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    76
\begin{quote}\index{generalizing induction formulae}%
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    77
\emph{Generalize goals for induction by replacing constants by variables.}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
\end{quote}
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    79
Of course one cannot do this na\"{\i}vely: @{term"itrev xs ys = rev xs"} is
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    80
just not true.  The correct generalization is
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
    81
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
(*<*)oops;(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
lemma "itrev xs ys = rev xs @ ys";
10362
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 9844
diff changeset
    84
(*<*)apply(induct_tac xs, simp_all)(*>*)
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    85
txt{*\noindent
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    86
If @{term"ys"} is replaced by @{term"[]"}, the right-hand side simplifies to
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    87
@{term"rev xs"}, as required.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    88
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    89
In this instance it was easy to guess the right generalization.
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    90
Other situations can require a good deal of creativity.  
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    91
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    92
Although we now have two variables, only @{term"xs"} is suitable for
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
    93
induction, and we repeat our proof attempt. Unfortunately, we are still
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    94
not there:
10362
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 9844
diff changeset
    95
@{subgoals[display,indent=0,goals_limit=1]}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    96
The induction hypothesis is still too weak, but this time it takes no
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    97
intuition to generalize: the problem is that @{term"ys"} is fixed throughout
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    98
the subgoal, but the induction hypothesis needs to be applied with
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
    99
@{term"a # ys"} instead of @{term"ys"}. Hence we prove the theorem
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
   100
for all @{term"ys"} instead of a fixed one:
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8745
diff changeset
   101
*};
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   102
(*<*)oops;(*>*)
10362
c6b197ccf1f1 *** empty log message ***
nipkow
parents: 9844
diff changeset
   103
lemma "\<forall>ys. itrev xs ys = rev xs @ ys";
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   104
(*<*)
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   105
by(induct_tac xs, simp_all);
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   106
(*>*)
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   107
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   108
text{*\noindent
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
   109
This time induction on @{term"xs"} followed by simplification succeeds. This
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   110
leads to another heuristic for generalization:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   111
\begin{quote}
9754
a123a64cadeb *** empty log message ***
nipkow
parents: 9723
diff changeset
   112
\emph{Generalize goals for induction by universally quantifying all free
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   113
variables {\em(except the induction variable itself!)}.}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   114
\end{quote}
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   115
This prevents trivial failures like the one above and does not affect the
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   116
validity of the goal.  However, this heuristic should not be applied blindly.
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   117
It is not always required, and the additional quantifiers can complicate
13081
ab4a3aef3591 *** empty log message ***
nipkow
parents: 11458
diff changeset
   118
matters in some cases. The variables that should be quantified are typically
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   119
those that change in recursive calls.
9644
6b0b6b471855 *** empty log message ***
nipkow
parents: 9541
diff changeset
   120
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   121
A final point worth mentioning is the orientation of the equation we just
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   122
proved: the more complex notion (@{term itrev}) is on the left-hand
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   123
side, the simpler one (@{term rev}) on the right-hand side. This constitutes
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   124
another, albeit weak heuristic that is not restricted to induction:
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   125
\begin{quote}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   126
  \emph{The right-hand side of an equation should (in some sense) be simpler
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   127
    than the left-hand side.}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   128
\end{quote}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   129
This heuristic is tricky to apply because it is not obvious that
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   130
@{term"rev xs @ ys"} is simpler than @{term"itrev xs ys"}. But see what
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   131
happens if you try to prove @{prop"rev xs @ ys = itrev xs ys"}!
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
   132
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   133
If you have tried these heuristics and still find your
10971
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
   134
induction does not go through, and no obvious lemma suggests itself, you may
6852682eaf16 *** empty log message ***
nipkow
parents: 10885
diff changeset
   135
need to generalize your proposition even further. This requires insight into
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   136
the problem at hand and is beyond simple rules of thumb.  
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   137
Additionally, you can read \S\ref{sec:advanced-ind}
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   138
to learn about some advanced techniques for inductive proofs.%
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 10971
diff changeset
   139
\index{induction heuristics|)}
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
   140
*}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   141
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   142
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   143
(*>*)