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