doc-src/TutorialI/ToyList/ToyList.thy
author nipkow
Tue, 25 Apr 2000 08:09:10 +0200
changeset 8771 026f37a86ea7
parent 8745 13b32661dde4
child 9458 c613cd06d5cf
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
theory ToyList = PreList:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
text{*\noindent
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
     4
HOL already has a predefined theory of lists called \isa{List} ---
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
     5
\isa{ToyList} is merely a small fragment of it chosen as an example. In
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     6
contrast to what is recommended in \S\ref{sec:Basic:Theories},
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
     7
\isa{ToyList} is not based on \isa{Main} but on \isa{PreList}, a
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     8
theory that contains pretty much everything but lists, thus avoiding
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     9
ambiguities caused by defining lists twice.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    10
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
datatype 'a list = Nil                          ("[]")
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
                 | Cons 'a "'a list"            (infixr "#" 65);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
The datatype\index{*datatype} \isaindexbold{list} introduces two
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
constructors \isaindexbold{Nil} and \isaindexbold{Cons}, the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
empty list and the operator that adds an element to the front of a list. For
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
example, the term \isa{Cons True (Cons   False Nil)} is a value of type
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
\isa{bool~list}, namely the list with the elements \isa{True} and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
\isa{False}. Because this notation becomes unwieldy very quickly, the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
datatype declaration is annotated with an alternative syntax: instead of
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
\isa{Nil} and \isa{Cons~$x$~$xs$} we can write
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
\isa{[]}\index{$HOL2list@\texttt{[]}|bold} and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
\isa{$x$~\#~$xs$}\index{$HOL2list@\texttt{\#}|bold}. In fact, this
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
alternative syntax is the standard syntax. Thus the list \isa{Cons True
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
(Cons False Nil)} becomes \isa{True \# False \# []}. The annotation
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
\isacommand{infixr}\indexbold{*infixr} means that \isa{\#} associates to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
the right, i.e.\ the term \isa{$x$ \# $y$ \# $z$} is read as \isa{$x$
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
\# ($y$ \# $z$)} and not as \isa{($x$ \# $y$) \# $z$}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    31
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
\begin{warn}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
  Syntax annotations are a powerful but completely optional feature. You
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    34
  could drop them from theory \isa{ToyList} and go back to the identifiers
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
  \isa{Nil} and \isa{Cons}. However, lists are such a central datatype
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    36
  that their syntax is highly customized. We recommend that novices should
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    37
  not use syntax annotations in their own theories.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
\end{warn}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    39
Next, two functions \isa{app} and \isaindexbold{rev} are declared:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
consts app :: "'a list \\<Rightarrow> 'a list \\<Rightarrow> 'a list"   (infixr "@" 65)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    43
       rev :: "'a list \\<Rightarrow> 'a list";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    46
\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    47
In contrast to ML, Isabelle insists on explicit declarations of all functions
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
(keyword \isacommand{consts}).  (Apart from the declaration-before-use
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
restriction, the order of items in a theory file is unconstrained.) Function
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    50
\isa{app} is annotated with concrete syntax too. Instead of the prefix
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    51
syntax \isa{app~$xs$~$ys$} the infix
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    52
\isa{$xs$~\at~$ys$}\index{$HOL2list@\texttt{\at}|bold} becomes the preferred
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    53
form. Both functions are defined recursively:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    54
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    55
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    56
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    57
"[] @ ys       = ys"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    58
"(x # xs) @ ys = x # (xs @ ys)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    59
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    60
primrec
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    61
"rev []        = []"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    62
"rev (x # xs)  = (rev xs) @ (x # [])";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    63
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    64
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    65
\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    66
The equations for \isa{app} and \isa{rev} hardly need comments:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    67
\isa{app} appends two lists and \isa{rev} reverses a list.  The keyword
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    68
\isacommand{primrec}\index{*primrec} indicates that the recursion is of a
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    69
particularly primitive kind where each recursive call peels off a datatype
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    70
constructor from one of the arguments.  Thus the
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    71
recursion always terminates, i.e.\ the function is \bfindex{total}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    72
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    73
The termination requirement is absolutely essential in HOL, a logic of total
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    74
functions. If we were to drop it, inconsistencies would quickly arise: the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    75
``definition'' $f(n) = f(n)+1$ immediately leads to $0 = 1$ by subtracting
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    76
$f(n)$ on both sides.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    77
% However, this is a subtle issue that we cannot discuss here further.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    78
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    79
\begin{warn}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    80
  As we have indicated, the desire for total functions is not a gratuitously
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    81
  imposed restriction but an essential characteristic of HOL. It is only
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    82
  because of totality that reasoning in HOL is comparatively easy.  More
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    83
  generally, the philosophy in HOL is not to allow arbitrary axioms (such as
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    84
  function definitions whose totality has not been proved) because they
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    85
  quickly lead to inconsistencies. Instead, fixed constructs for introducing
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    86
  types and functions are offered (such as \isacommand{datatype} and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    87
  \isacommand{primrec}) which are guaranteed to preserve consistency.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    88
\end{warn}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    89
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    90
A remark about syntax.  The textual definition of a theory follows a fixed
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    91
syntax with keywords like \isacommand{datatype} and \isacommand{end} (see
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    92
Fig.~\ref{fig:keywords} in Appendix~\ref{sec:Appendix} for a full list).
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    93
Embedded in this syntax are the types and formulae of HOL, whose syntax is
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    94
extensible, e.g.\ by new user-defined infix operators
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    95
(see~\ref{sec:infix-syntax}). To distinguish the two levels, everything
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    96
HOL-specific (terms and types) should be enclosed in
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    97
\texttt{"}\dots\texttt{"}. 
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    98
To lessen this burden, quotation marks around a single identifier can be
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    99
dropped, unless the identifier happens to be a keyword, as in
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   100
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   101
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   102
consts "end" :: "'a list \\<Rightarrow> 'a"
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   103
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   104
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   105
When Isabelle prints a syntax error message, it refers to the HOL syntax as
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   106
the \bfindex{inner syntax} and the enclosing theory language as the \bfindex{outer syntax}.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   107
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   108
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   109
\section{An introductory proof}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   110
\label{sec:intro-proof}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   111
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   112
Assuming you have input the declarations and definitions of \texttt{ToyList}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   113
presented so far, we are ready to prove a few simple theorems. This will
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   114
illustrate not just the basic proof commands but also the typical proof
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   115
process.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   116
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   117
\subsubsection*{Main goal: \texttt{rev(rev xs) = xs}}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   118
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   119
Our goal is to show that reversing a list twice produces the original
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   120
list. The input line
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   121
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   122
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   123
theorem rev_rev [simp]: "rev(rev xs) = xs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   124
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   125
txt{*\index{*theorem|bold}\index{*simp (attribute)|bold}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   126
\begin{itemize}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   127
\item
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   128
establishes a new theorem to be proved, namely \isa{rev(rev xs) = xs},
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   129
\item
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   130
gives that theorem the name \isa{rev_rev} by which it can be referred to,
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   131
\item
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   132
and tells Isabelle (via \isa{[simp]}) to use the theorem (once it has been
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   133
proved) as a simplification rule, i.e.\ all future proofs involving
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   134
simplification will replace occurrences of \isa{rev(rev xs)} by
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   135
\isa{xs}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   136
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   137
The name and the simplification attribute are optional.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   138
\end{itemize}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   139
Isabelle's response is to print
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   140
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   141
proof(prove):~step~0\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   142
\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   143
goal~(theorem~rev\_rev):\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   144
rev~(rev~xs)~=~xs\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   145
~1.~rev~(rev~xs)~=~xs
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   146
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   147
The first three lines tell us that we are 0 steps into the proof of
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   148
theorem \isa{rev_rev}; for compactness reasons we rarely show these
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   149
initial lines in this tutorial. The remaining lines display the current
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   150
proof state.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   151
Until we have finished a proof, the proof state always looks like this:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   152
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   153
$G$\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   154
~1.~$G\sb{1}$\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   155
~~\vdots~~\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   156
~$n$.~$G\sb{n}$
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   157
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   158
where $G$
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   159
is the overall goal that we are trying to prove, and the numbered lines
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   160
contain the subgoals $G\sb{1}$, \dots, $G\sb{n}$ that we need to prove to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   161
establish $G$. At \isa{step 0} there is only one subgoal, which is
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   162
identical with the overall goal.  Normally $G$ is constant and only serves as
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   163
a reminder. Hence we rarely show it in this tutorial.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   164
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   165
Let us now get back to \isa{rev(rev xs) = xs}. Properties of recursively
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   166
defined functions are best established by induction. In this case there is
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   167
not much choice except to induct on \isa{xs}:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   168
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   169
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   170
apply(induct_tac xs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   171
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   172
txt{*\noindent\index{*induct_tac}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   173
This tells Isabelle to perform induction on variable \isa{xs}. The suffix
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   174
\isa{tac} stands for ``tactic'', a synonym for ``theorem proving function''.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   175
By default, induction acts on the first subgoal. The new proof state contains
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   176
two subgoals, namely the base case (\isa{Nil}) and the induction step
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   177
(\isa{Cons}):
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   178
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   179
~1.~rev~(rev~[])~=~[]\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   180
~2.~{\isasymAnd}a~list.~rev(rev~list)~=~list~{\isasymLongrightarrow}~rev(rev(a~\#~list))~=~a~\#~list%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   181
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   182
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   183
The induction step is an example of the general format of a subgoal:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   184
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   185
~$i$.~{\indexboldpos{\isasymAnd}{$IsaAnd}}$x\sb{1}$~\dots~$x\sb{n}$.~{\it assumptions}~{\isasymLongrightarrow}~{\it conclusion}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   186
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   187
The prefix of bound variables \isasymAnd$x\sb{1}$~\dots~$x\sb{n}$ can be
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   188
ignored most of the time, or simply treated as a list of variables local to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   189
this subgoal. Their deeper significance is explained in \S\ref{sec:PCproofs}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   190
The {\it assumptions} are the local assumptions for this subgoal and {\it
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   191
  conclusion} is the actual proposition to be proved. Typical proof steps
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   192
that add new assumptions are induction or case distinction. In our example
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   193
the only assumption is the induction hypothesis \isa{rev (rev list) =
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   194
  list}, where \isa{list} is a variable name chosen by Isabelle. If there
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   195
are multiple assumptions, they are enclosed in the bracket pair
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   196
\indexboldpos{\isasymlbrakk}{$Isabrl} and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   197
\indexboldpos{\isasymrbrakk}{$Isabrr} and separated by semicolons.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   198
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   199
%FIXME indent!
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   200
Let us try to solve both goals automatically:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   201
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   202
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   203
apply(auto);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   204
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   205
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   206
This command tells Isabelle to apply a proof strategy called
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   207
\isa{auto} to all subgoals. Essentially, \isa{auto} tries to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   208
``simplify'' the subgoals.  In our case, subgoal~1 is solved completely (thanks
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   209
to the equation \isa{rev [] = []}) and disappears; the simplified version
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   210
of subgoal~2 becomes the new subgoal~1:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   211
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   212
~1.~\dots~rev(rev~list)~=~list~{\isasymLongrightarrow}~rev(rev~list~@~a~\#~[])~=~a~\#~list
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   213
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   214
In order to simplify this subgoal further, a lemma suggests itself.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   215
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   216
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   217
oops
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   218
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   219
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   220
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   221
\subsubsection*{First lemma: \texttt{rev(xs \at~ys) = (rev ys) \at~(rev xs)}}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   222
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   223
After abandoning the above proof attempt\indexbold{abandon proof} (at the shell level type
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   224
\isacommand{oops}) we start a new proof:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   225
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   226
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   227
lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   228
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   229
txt{*\noindent The keywords \isacommand{theorem}\index{*theorem} and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   230
\isacommand{lemma}\indexbold{*lemma} are interchangable and merely indicate
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   231
the importance we attach to a proposition. In general, we use the words
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   232
\emph{theorem}\index{theorem} and \emph{lemma}\index{lemma} pretty much
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   233
interchangeably.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   234
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   235
There are two variables that we could induct on: \isa{xs} and
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   236
\isa{ys}. Because \isa{\at} is defined by recursion on
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   237
the first argument, \isa{xs} is the correct one:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   238
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   239
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   240
apply(induct_tac xs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   241
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   242
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   243
This time not even the base case is solved automatically:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   244
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   245
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   246
apply(auto);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   247
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   248
txt{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   249
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   250
~1.~rev~ys~=~rev~ys~@~[]\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   251
~2. \dots
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   252
\end{isabellepar}%
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   253
Again, we need to abandon this proof attempt and prove another simple lemma first.
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   254
In the future the step of abandoning an incomplete proof before embarking on
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
   255
the proof of a lemma usually remains implicit.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   256
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   257
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   258
oops
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   259
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   260
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   261
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   262
\subsubsection*{Second lemma: \texttt{xs \at~[] = xs}}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   263
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   264
This time the canonical proof procedure
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   265
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   266
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   267
lemma app_Nil2 [simp]: "xs @ [] = xs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   268
apply(induct_tac xs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   269
apply(auto);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   270
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   271
txt{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   272
\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   273
leads to the desired message \isa{No subgoals!}:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   274
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   275
xs~@~[]~=~xs\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   276
No~subgoals!
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   277
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   278
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   279
We still need to confirm that the proof is now finished:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   280
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   281
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   282
.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   283
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   284
text{*\noindent\indexbold{$Isar@\texttt{.}}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   285
As a result of that final dot, Isabelle associates the lemma
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   286
just proved with its name. Notice that in the lemma \isa{app_Nil2} (as
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   287
printed out after the final dot) the free variable \isa{xs} has been
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   288
replaced by the unknown \isa{?xs}, just as explained in
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   289
\S\ref{sec:variables}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   290
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   291
Going back to the proof of the first lemma
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   292
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   293
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   294
lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   295
apply(induct_tac xs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   296
apply(auto);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   297
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   298
txt{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   299
\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   300
we find that this time \isa{auto} solves the base case, but the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   301
induction step merely simplifies to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   302
\begin{isabellepar}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   303
~1.~{\isasymAnd}a~list.\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   304
~~~~~~~rev~(list~@~ys)~=~rev~ys~@~rev~list~{\isasymLongrightarrow}\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   305
~~~~~~~(rev~ys~@~rev~list)~@~a~\#~[]~=~rev~ys~@~rev~list~@~a~\#~[]
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   306
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   307
Now we need to remember that \isa{\at} associates to the right, and that
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   308
\isa{\#} and \isa{\at} have the same priority (namely the \isa{65}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   309
in their \isacommand{infixr} annotation). Thus the conclusion really is
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   310
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   311
~~~~~(rev~ys~@~rev~list)~@~(a~\#~[])~=~rev~ys~@~(rev~list~@~(a~\#~[]))%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   312
\end{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   313
and the missing lemma is associativity of \isa{\at}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   314
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   315
\subsubsection*{Third lemma: \texttt{(xs \at~ys) \at~zs = xs \at~(ys \at~zs)}}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   316
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   317
Abandoning the previous proof, the canonical proof procedure
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   318
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   319
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   320
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   321
txt_raw{*\begin{comment}*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   322
oops
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   323
text_raw{*\end{comment}*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   324
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   325
lemma app_assoc [simp]: "(xs @ ys) @ zs = xs @ (ys @ zs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   326
apply(induct_tac xs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   327
apply(auto).;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   328
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   329
text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   330
\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   331
succeeds without further ado.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   332
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   333
Now we can go back and prove the first lemma
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   334
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   335
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   336
lemma rev_app [simp]: "rev(xs @ ys) = (rev ys) @ (rev xs)";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   337
apply(induct_tac xs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   338
apply(auto).;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   339
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   340
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   341
and then solve our main theorem:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   342
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   343
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   344
theorem rev_rev [simp]: "rev(rev xs) = xs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   345
apply(induct_tac xs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   346
apply(auto).;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   347
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   348
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   349
The final \isa{end} tells Isabelle to close the current theory because
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   350
we are finished with its development:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   351
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   352
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
   353
end