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