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