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