author | wenzelm |
Fri, 24 Aug 2012 20:47:33 +0200 | |
changeset 48924 | 27d8ccd1906c |
parent 48519 | 5deda0549f97 |
permissions | -rw-r--r-- |
9722 | 1 |
% |
2 |
\begin{isabellebody}% |
|
9924 | 3 |
\def\isabellecontext{ToyList}% |
17056 | 4 |
% |
5 |
\isadelimtheory |
|
6 |
% |
|
7 |
\endisadelimtheory |
|
8 |
% |
|
9 |
\isatagtheory |
|
17175 | 10 |
\isacommand{theory}\isamarkupfalse% |
11 |
\ ToyList\isanewline |
|
26839 | 12 |
\isakeyword{imports}\ Datatype\isanewline |
17056 | 13 |
\isakeyword{begin}% |
14 |
\endisatagtheory |
|
15 |
{\isafoldtheory}% |
|
16 |
% |
|
17 |
\isadelimtheory |
|
18 |
% |
|
19 |
\endisadelimtheory |
|
11866 | 20 |
% |
8749 | 21 |
\begin{isamarkuptext}% |
22 |
\noindent |
|
8771 | 23 |
HOL already has a predefined theory of lists called \isa{List} --- |
24 |
\isa{ToyList} is merely a small fragment of it chosen as an example. In |
|
8749 | 25 |
contrast to what is recommended in \S\ref{sec:Basic:Theories}, |
26839 | 26 |
\isa{ToyList} is not based on \isa{Main} but on \isa{Datatype}, a |
8749 | 27 |
theory that contains pretty much everything but lists, thus avoiding |
28 |
ambiguities caused by defining lists twice.% |
|
29 |
\end{isamarkuptext}% |
|
17175 | 30 |
\isamarkuptrue% |
31 |
\isacommand{datatype}\isamarkupfalse% |
|
40406 | 32 |
\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{3D}{\isacharequal}}\ Nil\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\isanewline |
33 |
\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{7C}{\isacharbar}}\ Cons\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequoteclose}}\ \ \ \ \ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixr}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{23}{\isacharhash}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{6}}{\isadigit{5}}{\isaliteral{29}{\isacharparenright}}% |
|
8749 | 34 |
\begin{isamarkuptext}% |
35 |
\noindent |
|
12327 | 36 |
The datatype\index{datatype@\isacommand {datatype} (command)} |
37 |
\tydx{list} introduces two |
|
11428 | 38 |
constructors \cdx{Nil} and \cdx{Cons}, the |
9541 | 39 |
empty~list and the operator that adds an element to the front of a list. For |
9792 | 40 |
example, the term \isa{Cons True (Cons False Nil)} is a value of |
41 |
type \isa{bool\ list}, namely the list with the elements \isa{True} and |
|
11450 | 42 |
\isa{False}. Because this notation quickly becomes unwieldy, the |
8749 | 43 |
datatype declaration is annotated with an alternative syntax: instead of |
9541 | 44 |
\isa{Nil} and \isa{Cons x xs} we can write |
40406 | 45 |
\isa{{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}}\index{$HOL2list@\isa{[]}|bold} and |
46 |
\isa{x\ {\isaliteral{23}{\isacharhash}}\ xs}\index{$HOL2list@\isa{\#}|bold}. In fact, this |
|
11450 | 47 |
alternative syntax is the familiar one. Thus the list \isa{Cons True |
40406 | 48 |
(Cons False Nil)} becomes \isa{True\ {\isaliteral{23}{\isacharhash}}\ False\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}}. The annotation |
11428 | 49 |
\isacommand{infixr}\index{infixr@\isacommand{infixr} (annotation)} |
40406 | 50 |
means that \isa{{\isaliteral{23}{\isacharhash}}} associates to |
51 |
the right: the term \isa{x\ {\isaliteral{23}{\isacharhash}}\ y\ {\isaliteral{23}{\isacharhash}}\ z} is read as \isa{x\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{23}{\isacharhash}}\ z{\isaliteral{29}{\isacharparenright}}} |
|
52 |
and not as \isa{{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{23}{\isacharhash}}\ z}. |
|
53 |
The \isa{{\isadigit{6}}{\isadigit{5}}} is the priority of the infix \isa{{\isaliteral{23}{\isacharhash}}}. |
|
8749 | 54 |
|
55 |
\begin{warn} |
|
13191 | 56 |
Syntax annotations can be powerful, but they are difficult to master and |
11456 | 57 |
are never necessary. You |
8771 | 58 |
could drop them from theory \isa{ToyList} and go back to the identifiers |
27015 | 59 |
\isa{Nil} and \isa{Cons}. Novices should avoid using |
10795 | 60 |
syntax annotations in their own theories. |
8749 | 61 |
\end{warn} |
27015 | 62 |
Next, two functions \isa{app} and \cdx{rev} are defined recursively, |
63 |
in this order, because Isabelle insists on definition before use:% |
|
8749 | 64 |
\end{isamarkuptext}% |
17175 | 65 |
\isamarkuptrue% |
66 |
\isacommand{primrec}\isamarkupfalse% |
|
40406 | 67 |
\ app\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixr}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{40}{\isacharat}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{6}}{\isadigit{5}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline |
68 |
{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{40}{\isacharat}}\ ys\ \ \ \ \ \ \ {\isaliteral{3D}{\isacharequal}}\ ys{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline |
|
69 |
{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{40}{\isacharat}}\ ys\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{28}{\isacharparenleft}}xs\ {\isaliteral{40}{\isacharat}}\ ys{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline |
|
17175 | 70 |
\isanewline |
71 |
\isacommand{primrec}\isamarkupfalse% |
|
40406 | 72 |
\ rev\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{where}\isanewline |
73 |
{\isaliteral{22}{\isachardoublequoteopen}}rev\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ \ \ \ \ \ \ \ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\isanewline |
|
74 |
{\isaliteral{22}{\isachardoublequoteopen}}rev\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ xs{\isaliteral{29}{\isacharparenright}}\ \ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% |
|
8749 | 75 |
\begin{isamarkuptext}% |
27015 | 76 |
\noindent |
77 |
Each function definition is of the form |
|
78 |
\begin{center} |
|
40406 | 79 |
\isacommand{primrec} \textit{name} \isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}} \textit{type} \textit{(optional syntax)} \isakeyword{where} \textit{equations} |
27015 | 80 |
\end{center} |
40406 | 81 |
The equations must be separated by \isa{{\isaliteral{7C}{\isacharbar}}}. |
27015 | 82 |
% |
83 |
Function \isa{app} is annotated with concrete syntax. Instead of the |
|
84 |
prefix syntax \isa{app\ xs\ ys} the infix |
|
40406 | 85 |
\isa{xs\ {\isaliteral{40}{\isacharat}}\ ys}\index{$HOL2list@\isa{\at}|bold} becomes the preferred |
27015 | 86 |
form. |
87 |
||
88 |
\index{*rev (constant)|(}\index{append function|(} |
|
10790 | 89 |
The equations for \isa{app} and \isa{rev} hardly need comments: |
90 |
\isa{app} appends two lists and \isa{rev} reverses a list. The |
|
11428 | 91 |
keyword \commdx{primrec} indicates that the recursion is |
10790 | 92 |
of a particularly primitive kind where each recursive call peels off a datatype |
8771 | 93 |
constructor from one of the arguments. Thus the |
10654 | 94 |
recursion always terminates, i.e.\ the function is \textbf{total}. |
11428 | 95 |
\index{functions!total} |
8749 | 96 |
|
97 |
The termination requirement is absolutely essential in HOL, a logic of total |
|
98 |
functions. If we were to drop it, inconsistencies would quickly arise: the |
|
99 |
``definition'' $f(n) = f(n)+1$ immediately leads to $0 = 1$ by subtracting |
|
100 |
$f(n)$ on both sides. |
|
101 |
% However, this is a subtle issue that we cannot discuss here further. |
|
102 |
||
103 |
\begin{warn} |
|
11456 | 104 |
As we have indicated, the requirement for total functions is an essential characteristic of HOL\@. It is only |
8749 | 105 |
because of totality that reasoning in HOL is comparatively easy. More |
11456 | 106 |
generally, the philosophy in HOL is to refrain from asserting arbitrary axioms (such as |
8749 | 107 |
function definitions whose totality has not been proved) because they |
108 |
quickly lead to inconsistencies. Instead, fixed constructs for introducing |
|
109 |
types and functions are offered (such as \isacommand{datatype} and |
|
110 |
\isacommand{primrec}) which are guaranteed to preserve consistency. |
|
111 |
\end{warn} |
|
112 |
||
11456 | 113 |
\index{syntax}% |
8749 | 114 |
A remark about syntax. The textual definition of a theory follows a fixed |
10971 | 115 |
syntax with keywords like \isacommand{datatype} and \isacommand{end}. |
116 |
% (see Fig.~\ref{fig:keywords} in Appendix~\ref{sec:Appendix} for a full list). |
|
8749 | 117 |
Embedded in this syntax are the types and formulae of HOL, whose syntax is |
12627 | 118 |
extensible (see \S\ref{sec:concrete-syntax}), e.g.\ by new user-defined infix operators. |
10971 | 119 |
To distinguish the two levels, everything |
8749 | 120 |
HOL-specific (terms and types) should be enclosed in |
121 |
\texttt{"}\dots\texttt{"}. |
|
122 |
To lessen this burden, quotation marks around a single identifier can be |
|
27015 | 123 |
dropped, unless the identifier happens to be a keyword, for example |
124 |
\isa{"end"}. |
|
8749 | 125 |
When Isabelle prints a syntax error message, it refers to the HOL syntax as |
11456 | 126 |
the \textbf{inner syntax} and the enclosing theory language as the \textbf{outer syntax}. |
8749 | 127 |
|
38430
254a021ed66e
tuned text about "value" and added note on comments.
nipkow
parents:
27015
diff
changeset
|
128 |
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
|
129 |
|
25342 | 130 |
\section{Evaluation} |
131 |
\index{evaluation} |
|
132 |
||
133 |
Assuming you have processed the declarations and definitions of |
|
134 |
\texttt{ToyList} presented so far, you may want to test your |
|
135 |
functions by running them. For example, what is the value of |
|
40406 | 136 |
\isa{rev\ {\isaliteral{28}{\isacharparenleft}}True\ {\isaliteral{23}{\isacharhash}}\ False\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}}? Command% |
25342 | 137 |
\end{isamarkuptext}% |
138 |
\isamarkuptrue% |
|
139 |
\isacommand{value}\isamarkupfalse% |
|
40406 | 140 |
\ {\isaliteral{22}{\isachardoublequoteopen}}rev\ {\isaliteral{28}{\isacharparenleft}}True\ {\isaliteral{23}{\isacharhash}}\ False\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% |
25342 | 141 |
\begin{isamarkuptext}% |
40406 | 142 |
\noindent yields the correct result \isa{False\ {\isaliteral{23}{\isacharhash}}\ True\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}}. |
25342 | 143 |
But we can go beyond mere functional programming and evaluate terms with |
144 |
variables in them, executing functions symbolically:% |
|
145 |
\end{isamarkuptext}% |
|
146 |
\isamarkuptrue% |
|
38430
254a021ed66e
tuned text about "value" and added note on comments.
nipkow
parents:
27015
diff
changeset
|
147 |
\isacommand{value}\isamarkupfalse% |
40406 | 148 |
\ {\isaliteral{22}{\isachardoublequoteopen}}rev\ {\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{23}{\isacharhash}}\ b\ {\isaliteral{23}{\isacharhash}}\ c\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% |
25342 | 149 |
\begin{isamarkuptext}% |
40406 | 150 |
\noindent yields \isa{c\ {\isaliteral{23}{\isacharhash}}\ b\ {\isaliteral{23}{\isacharhash}}\ a\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}}. |
38432
439f50a241c1
Using type real does not require a separate logic now.
nipkow
parents:
38430
diff
changeset
|
151 |
|
10878 | 152 |
\section{An Introductory Proof} |
8749 | 153 |
\label{sec:intro-proof} |
154 |
||
25342 | 155 |
Having convinced ourselves (as well as one can by testing) that our |
156 |
definitions capture our intentions, we are ready to prove a few simple |
|
16409 | 157 |
theorems. This will illustrate not just the basic proof commands but |
158 |
also the typical proof process. |
|
8749 | 159 |
|
11457 | 160 |
\subsubsection*{Main Goal.} |
8749 | 161 |
|
162 |
Our goal is to show that reversing a list twice produces the original |
|
11456 | 163 |
list.% |
8749 | 164 |
\end{isamarkuptext}% |
17175 | 165 |
\isamarkuptrue% |
166 |
\isacommand{theorem}\isamarkupfalse% |
|
40406 | 167 |
\ rev{\isaliteral{5F}{\isacharunderscore}}rev\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}rev{\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}% |
17056 | 168 |
\isadelimproof |
169 |
% |
|
170 |
\endisadelimproof |
|
171 |
% |
|
172 |
\isatagproof |
|
16069 | 173 |
% |
174 |
\begin{isamarkuptxt}% |
|
175 |
\index{theorem@\isacommand {theorem} (command)|bold}% |
|
176 |
\noindent |
|
177 |
This \isacommand{theorem} command does several things: |
|
178 |
\begin{itemize} |
|
179 |
\item |
|
40406 | 180 |
It establishes a new theorem to be proved, namely \isa{rev\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ xs}. |
16069 | 181 |
\item |
40406 | 182 |
It gives that theorem the name \isa{rev{\isaliteral{5F}{\isacharunderscore}}rev}, for later reference. |
16069 | 183 |
\item |
184 |
It tells Isabelle (via the bracketed attribute \attrdx{simp}) to take the eventual theorem as a simplification rule: future proofs involving |
|
40406 | 185 |
simplification will replace occurrences of \isa{rev\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}} by |
16069 | 186 |
\isa{xs}. |
187 |
\end{itemize} |
|
188 |
The name and the simplification attribute are optional. |
|
189 |
Isabelle's response is to print the initial proof state consisting |
|
190 |
of some header information (like how many subgoals there are) followed by |
|
191 |
\begin{isabelle}% |
|
40406 | 192 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ rev\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ xs% |
16069 | 193 |
\end{isabelle} |
194 |
For compactness reasons we omit the header in this tutorial. |
|
195 |
Until we have finished a proof, the \rmindex{proof state} proper |
|
196 |
always looks like this: |
|
197 |
\begin{isabelle} |
|
198 |
~1.~$G\sb{1}$\isanewline |
|
199 |
~~\vdots~~\isanewline |
|
200 |
~$n$.~$G\sb{n}$ |
|
201 |
\end{isabelle} |
|
202 |
The numbered lines contain the subgoals $G\sb{1}$, \dots, $G\sb{n}$ |
|
203 |
that we need to prove to establish the main goal.\index{subgoals} |
|
204 |
Initially there is only one subgoal, which is identical with the |
|
205 |
main goal. (If you always want to see the main goal as well, |
|
206 |
set the flag \isa{Proof.show_main_goal}\index{*show_main_goal (flag)} |
|
207 |
--- this flag used to be set by default.) |
|
208 |
||
40406 | 209 |
Let us now get back to \isa{rev\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ xs}. Properties of recursively |
16069 | 210 |
defined functions are best established by induction. In this case there is |
211 |
nothing obvious except induction on \isa{xs}:% |
|
212 |
\end{isamarkuptxt}% |
|
17175 | 213 |
\isamarkuptrue% |
214 |
\isacommand{apply}\isamarkupfalse% |
|
40406 | 215 |
{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ xs{\isaliteral{29}{\isacharparenright}}% |
16069 | 216 |
\begin{isamarkuptxt}% |
217 |
\noindent\index{*induct_tac (method)}% |
|
218 |
This tells Isabelle to perform induction on variable \isa{xs}. The suffix |
|
219 |
\isa{tac} stands for \textbf{tactic},\index{tactics} |
|
220 |
a synonym for ``theorem proving function''. |
|
221 |
By default, induction acts on the first subgoal. The new proof state contains |
|
222 |
two subgoals, namely the base case (\isa{Nil}) and the induction step |
|
223 |
(\isa{Cons}): |
|
224 |
\begin{isabelle}% |
|
40406 | 225 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ rev\ {\isaliteral{28}{\isacharparenleft}}rev\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\isanewline |
226 |
\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}a\ list{\isaliteral{2E}{\isachardot}}\isanewline |
|
227 |
\isaindent{\ {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ \ \ \ }rev\ {\isaliteral{28}{\isacharparenleft}}rev\ list{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ list\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ rev\ {\isaliteral{28}{\isacharparenleft}}rev\ {\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{23}{\isacharhash}}\ list{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ a\ {\isaliteral{23}{\isacharhash}}\ list% |
|
16069 | 228 |
\end{isabelle} |
229 |
||
230 |
The induction step is an example of the general format of a subgoal:\index{subgoals} |
|
231 |
\begin{isabelle} |
|
232 |
~$i$.~{\isasymAnd}$x\sb{1}$~\dots$x\sb{n}$.~{\it assumptions}~{\isasymLongrightarrow}~{\it conclusion} |
|
233 |
\end{isabelle}\index{$IsaAnd@\isasymAnd|bold} |
|
234 |
The prefix of bound variables \isasymAnd$x\sb{1}$~\dots~$x\sb{n}$ can be |
|
235 |
ignored most of the time, or simply treated as a list of variables local to |
|
236 |
this subgoal. Their deeper significance is explained in Chapter~\ref{chap:rules}. |
|
237 |
The {\it assumptions}\index{assumptions!of subgoal} |
|
238 |
are the local assumptions for this subgoal and {\it |
|
239 |
conclusion}\index{conclusion!of subgoal} is the actual proposition to be proved. |
|
240 |
Typical proof steps |
|
241 |
that add new assumptions are induction and case distinction. In our example |
|
40406 | 242 |
the only assumption is the induction hypothesis \isa{rev\ {\isaliteral{28}{\isacharparenleft}}rev\ list{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ list}, where \isa{list} is a variable name chosen by Isabelle. If there |
16069 | 243 |
are multiple assumptions, they are enclosed in the bracket pair |
244 |
\indexboldpos{\isasymlbrakk}{$Isabrl} and |
|
245 |
\indexboldpos{\isasymrbrakk}{$Isabrr} and separated by semicolons. |
|
246 |
||
247 |
Let us try to solve both goals automatically:% |
|
248 |
\end{isamarkuptxt}% |
|
17175 | 249 |
\isamarkuptrue% |
250 |
\isacommand{apply}\isamarkupfalse% |
|
40406 | 251 |
{\isaliteral{28}{\isacharparenleft}}auto{\isaliteral{29}{\isacharparenright}}% |
16069 | 252 |
\begin{isamarkuptxt}% |
253 |
\noindent |
|
254 |
This command tells Isabelle to apply a proof strategy called |
|
255 |
\isa{auto} to all subgoals. Essentially, \isa{auto} tries to |
|
256 |
simplify the subgoals. In our case, subgoal~1 is solved completely (thanks |
|
40406 | 257 |
to the equation \isa{rev\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}}) and disappears; the simplified version |
16069 | 258 |
of subgoal~2 becomes the new subgoal~1: |
259 |
\begin{isabelle}% |
|
40406 | 260 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}a\ list{\isaliteral{2E}{\isachardot}}\isanewline |
261 |
\isaindent{\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ \ \ \ }rev\ {\isaliteral{28}{\isacharparenleft}}rev\ list{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ list\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ rev\ {\isaliteral{28}{\isacharparenleft}}rev\ list\ {\isaliteral{40}{\isacharat}}\ a\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ a\ {\isaliteral{23}{\isacharhash}}\ list% |
|
16069 | 262 |
\end{isabelle} |
263 |
In order to simplify this subgoal further, a lemma suggests itself.% |
|
264 |
\end{isamarkuptxt}% |
|
17175 | 265 |
\isamarkuptrue% |
17056 | 266 |
% |
267 |
\endisatagproof |
|
268 |
{\isafoldproof}% |
|
269 |
% |
|
270 |
\isadelimproof |
|
271 |
% |
|
272 |
\endisadelimproof |
|
8749 | 273 |
% |
11428 | 274 |
\isamarkupsubsubsection{First Lemma% |
10395 | 275 |
} |
11866 | 276 |
\isamarkuptrue% |
9723 | 277 |
% |
8749 | 278 |
\begin{isamarkuptext}% |
11428 | 279 |
\indexbold{abandoning a proof}\indexbold{proofs!abandoning} |
280 |
After abandoning the above proof attempt (at the shell level type |
|
281 |
\commdx{oops}) we start a new proof:% |
|
8749 | 282 |
\end{isamarkuptext}% |
17175 | 283 |
\isamarkuptrue% |
284 |
\isacommand{lemma}\isamarkupfalse% |
|
40406 | 285 |
\ rev{\isaliteral{5F}{\isacharunderscore}}app\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}rev{\isaliteral{28}{\isacharparenleft}}xs\ {\isaliteral{40}{\isacharat}}\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}rev\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% |
17056 | 286 |
\isadelimproof |
287 |
% |
|
288 |
\endisadelimproof |
|
289 |
% |
|
290 |
\isatagproof |
|
16069 | 291 |
% |
292 |
\begin{isamarkuptxt}% |
|
293 |
\noindent The keywords \commdx{theorem} and |
|
294 |
\commdx{lemma} are interchangeable and merely indicate |
|
295 |
the importance we attach to a proposition. Therefore we use the words |
|
296 |
\emph{theorem} and \emph{lemma} pretty much interchangeably, too. |
|
297 |
||
298 |
There are two variables that we could induct on: \isa{xs} and |
|
40406 | 299 |
\isa{ys}. Because \isa{{\isaliteral{40}{\isacharat}}} is defined by recursion on |
16069 | 300 |
the first argument, \isa{xs} is the correct one:% |
301 |
\end{isamarkuptxt}% |
|
17175 | 302 |
\isamarkuptrue% |
303 |
\isacommand{apply}\isamarkupfalse% |
|
40406 | 304 |
{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ xs{\isaliteral{29}{\isacharparenright}}% |
16069 | 305 |
\begin{isamarkuptxt}% |
306 |
\noindent |
|
307 |
This time not even the base case is solved automatically:% |
|
308 |
\end{isamarkuptxt}% |
|
17175 | 309 |
\isamarkuptrue% |
310 |
\isacommand{apply}\isamarkupfalse% |
|
40406 | 311 |
{\isaliteral{28}{\isacharparenleft}}auto{\isaliteral{29}{\isacharparenright}}% |
16069 | 312 |
\begin{isamarkuptxt}% |
313 |
\begin{isabelle}% |
|
40406 | 314 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ rev\ ys\ {\isaliteral{3D}{\isacharequal}}\ rev\ ys\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}% |
16069 | 315 |
\end{isabelle} |
316 |
Again, we need to abandon this proof attempt and prove another simple lemma |
|
317 |
first. In the future the step of abandoning an incomplete proof before |
|
318 |
embarking on the proof of a lemma usually remains implicit.% |
|
319 |
\end{isamarkuptxt}% |
|
17175 | 320 |
\isamarkuptrue% |
17056 | 321 |
% |
322 |
\endisatagproof |
|
323 |
{\isafoldproof}% |
|
324 |
% |
|
325 |
\isadelimproof |
|
326 |
% |
|
327 |
\endisadelimproof |
|
8749 | 328 |
% |
11428 | 329 |
\isamarkupsubsubsection{Second Lemma% |
10395 | 330 |
} |
11866 | 331 |
\isamarkuptrue% |
9723 | 332 |
% |
8749 | 333 |
\begin{isamarkuptext}% |
11456 | 334 |
We again try the canonical proof procedure:% |
8749 | 335 |
\end{isamarkuptext}% |
17175 | 336 |
\isamarkuptrue% |
337 |
\isacommand{lemma}\isamarkupfalse% |
|
40406 | 338 |
\ app{\isaliteral{5F}{\isacharunderscore}}Nil{\isadigit{2}}\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}xs\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline |
17056 | 339 |
% |
340 |
\isadelimproof |
|
341 |
% |
|
342 |
\endisadelimproof |
|
343 |
% |
|
344 |
\isatagproof |
|
17175 | 345 |
\isacommand{apply}\isamarkupfalse% |
40406 | 346 |
{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ xs{\isaliteral{29}{\isacharparenright}}\isanewline |
17175 | 347 |
\isacommand{apply}\isamarkupfalse% |
40406 | 348 |
{\isaliteral{28}{\isacharparenleft}}auto{\isaliteral{29}{\isacharparenright}}% |
16069 | 349 |
\begin{isamarkuptxt}% |
350 |
\noindent |
|
40406 | 351 |
It works, yielding the desired message \isa{No\ subgoals{\isaliteral{21}{\isacharbang}}}: |
16069 | 352 |
\begin{isabelle}% |
40406 | 353 |
xs\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ xs\isanewline |
354 |
No\ subgoals{\isaliteral{21}{\isacharbang}}% |
|
16069 | 355 |
\end{isabelle} |
356 |
We still need to confirm that the proof is now finished:% |
|
357 |
\end{isamarkuptxt}% |
|
17175 | 358 |
\isamarkuptrue% |
359 |
\isacommand{done}\isamarkupfalse% |
|
360 |
% |
|
17056 | 361 |
\endisatagproof |
362 |
{\isafoldproof}% |
|
363 |
% |
|
364 |
\isadelimproof |
|
365 |
% |
|
366 |
\endisadelimproof |
|
11866 | 367 |
% |
8749 | 368 |
\begin{isamarkuptext}% |
11428 | 369 |
\noindent |
370 |
As a result of that final \commdx{done}, Isabelle associates the lemma just proved |
|
10171 | 371 |
with its name. In this tutorial, we sometimes omit to show that final \isacommand{done} |
372 |
if it is obvious from the context that the proof is finished. |
|
373 |
||
374 |
% Instead of \isacommand{apply} followed by a dot, you can simply write |
|
375 |
% \isacommand{by}\indexbold{by}, which we do most of the time. |
|
40406 | 376 |
Notice that in lemma \isa{app{\isaliteral{5F}{\isacharunderscore}}Nil{\isadigit{2}}}, |
10971 | 377 |
as printed out after the final \isacommand{done}, the free variable \isa{xs} has been |
40406 | 378 |
replaced by the unknown \isa{{\isaliteral{3F}{\isacharquery}}xs}, just as explained in |
9792 | 379 |
\S\ref{sec:variables}. |
8749 | 380 |
|
381 |
Going back to the proof of the first lemma% |
|
382 |
\end{isamarkuptext}% |
|
17175 | 383 |
\isamarkuptrue% |
384 |
\isacommand{lemma}\isamarkupfalse% |
|
40406 | 385 |
\ rev{\isaliteral{5F}{\isacharunderscore}}app\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}rev{\isaliteral{28}{\isacharparenleft}}xs\ {\isaliteral{40}{\isacharat}}\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}rev\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline |
17056 | 386 |
% |
387 |
\isadelimproof |
|
388 |
% |
|
389 |
\endisadelimproof |
|
390 |
% |
|
391 |
\isatagproof |
|
17175 | 392 |
\isacommand{apply}\isamarkupfalse% |
40406 | 393 |
{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ xs{\isaliteral{29}{\isacharparenright}}\isanewline |
17175 | 394 |
\isacommand{apply}\isamarkupfalse% |
40406 | 395 |
{\isaliteral{28}{\isacharparenleft}}auto{\isaliteral{29}{\isacharparenright}}% |
16069 | 396 |
\begin{isamarkuptxt}% |
397 |
\noindent |
|
398 |
we find that this time \isa{auto} solves the base case, but the |
|
399 |
induction step merely simplifies to |
|
400 |
\begin{isabelle}% |
|
40406 | 401 |
\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}a\ list{\isaliteral{2E}{\isachardot}}\isanewline |
402 |
\isaindent{\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ \ \ \ }rev\ {\isaliteral{28}{\isacharparenleft}}list\ {\isaliteral{40}{\isacharat}}\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ rev\ ys\ {\isaliteral{40}{\isacharat}}\ rev\ list\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline |
|
403 |
\isaindent{\ {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ \ \ \ }{\isaliteral{28}{\isacharparenleft}}rev\ ys\ {\isaliteral{40}{\isacharat}}\ rev\ list{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{40}{\isacharat}}\ a\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ rev\ ys\ {\isaliteral{40}{\isacharat}}\ rev\ list\ {\isaliteral{40}{\isacharat}}\ a\ {\isaliteral{23}{\isacharhash}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}% |
|
16069 | 404 |
\end{isabelle} |
40406 | 405 |
Now we need to remember that \isa{{\isaliteral{40}{\isacharat}}} associates to the right, and that |
406 |
\isa{{\isaliteral{23}{\isacharhash}}} and \isa{{\isaliteral{40}{\isacharat}}} have the same priority (namely the \isa{{\isadigit{6}}{\isadigit{5}}} |
|
16069 | 407 |
in their \isacommand{infixr} annotation). Thus the conclusion really is |
408 |
\begin{isabelle} |
|
409 |
~~~~~(rev~ys~@~rev~list)~@~(a~\#~[])~=~rev~ys~@~(rev~list~@~(a~\#~[])) |
|
410 |
\end{isabelle} |
|
40406 | 411 |
and the missing lemma is associativity of \isa{{\isaliteral{40}{\isacharat}}}.% |
16069 | 412 |
\end{isamarkuptxt}% |
17175 | 413 |
\isamarkuptrue% |
17056 | 414 |
% |
415 |
\endisatagproof |
|
416 |
{\isafoldproof}% |
|
417 |
% |
|
418 |
\isadelimproof |
|
419 |
% |
|
420 |
\endisadelimproof |
|
8749 | 421 |
% |
11456 | 422 |
\isamarkupsubsubsection{Third Lemma% |
10395 | 423 |
} |
11866 | 424 |
\isamarkuptrue% |
9723 | 425 |
% |
426 |
\begin{isamarkuptext}% |
|
11456 | 427 |
Abandoning the previous attempt, the canonical proof procedure |
428 |
succeeds without further ado.% |
|
9723 | 429 |
\end{isamarkuptext}% |
17175 | 430 |
\isamarkuptrue% |
431 |
\isacommand{lemma}\isamarkupfalse% |
|
40406 | 432 |
\ app{\isaliteral{5F}{\isacharunderscore}}assoc\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}xs\ {\isaliteral{40}{\isacharat}}\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{40}{\isacharat}}\ zs\ {\isaliteral{3D}{\isacharequal}}\ xs\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{28}{\isacharparenleft}}ys\ {\isaliteral{40}{\isacharat}}\ zs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline |
17056 | 433 |
% |
434 |
\isadelimproof |
|
435 |
% |
|
436 |
\endisadelimproof |
|
437 |
% |
|
438 |
\isatagproof |
|
17175 | 439 |
\isacommand{apply}\isamarkupfalse% |
40406 | 440 |
{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ xs{\isaliteral{29}{\isacharparenright}}\isanewline |
17175 | 441 |
\isacommand{apply}\isamarkupfalse% |
40406 | 442 |
{\isaliteral{28}{\isacharparenleft}}auto{\isaliteral{29}{\isacharparenright}}\isanewline |
17175 | 443 |
\isacommand{done}\isamarkupfalse% |
444 |
% |
|
17056 | 445 |
\endisatagproof |
446 |
{\isafoldproof}% |
|
447 |
% |
|
448 |
\isadelimproof |
|
449 |
% |
|
450 |
\endisadelimproof |
|
11866 | 451 |
% |
8749 | 452 |
\begin{isamarkuptext}% |
453 |
\noindent |
|
11456 | 454 |
Now we can prove the first lemma:% |
8749 | 455 |
\end{isamarkuptext}% |
17175 | 456 |
\isamarkuptrue% |
457 |
\isacommand{lemma}\isamarkupfalse% |
|
40406 | 458 |
\ rev{\isaliteral{5F}{\isacharunderscore}}app\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}rev{\isaliteral{28}{\isacharparenleft}}xs\ {\isaliteral{40}{\isacharat}}\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}rev\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{40}{\isacharat}}\ {\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline |
17056 | 459 |
% |
460 |
\isadelimproof |
|
461 |
% |
|
462 |
\endisadelimproof |
|
463 |
% |
|
464 |
\isatagproof |
|
17175 | 465 |
\isacommand{apply}\isamarkupfalse% |
40406 | 466 |
{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ xs{\isaliteral{29}{\isacharparenright}}\isanewline |
17175 | 467 |
\isacommand{apply}\isamarkupfalse% |
40406 | 468 |
{\isaliteral{28}{\isacharparenleft}}auto{\isaliteral{29}{\isacharparenright}}\isanewline |
17175 | 469 |
\isacommand{done}\isamarkupfalse% |
470 |
% |
|
17056 | 471 |
\endisatagproof |
472 |
{\isafoldproof}% |
|
473 |
% |
|
474 |
\isadelimproof |
|
475 |
% |
|
476 |
\endisadelimproof |
|
11866 | 477 |
% |
8749 | 478 |
\begin{isamarkuptext}% |
479 |
\noindent |
|
11456 | 480 |
Finally, we prove our main theorem:% |
8749 | 481 |
\end{isamarkuptext}% |
17175 | 482 |
\isamarkuptrue% |
483 |
\isacommand{theorem}\isamarkupfalse% |
|
40406 | 484 |
\ rev{\isaliteral{5F}{\isacharunderscore}}rev\ {\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}rev{\isaliteral{28}{\isacharparenleft}}rev\ xs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline |
17056 | 485 |
% |
486 |
\isadelimproof |
|
487 |
% |
|
488 |
\endisadelimproof |
|
489 |
% |
|
490 |
\isatagproof |
|
17175 | 491 |
\isacommand{apply}\isamarkupfalse% |
40406 | 492 |
{\isaliteral{28}{\isacharparenleft}}induct{\isaliteral{5F}{\isacharunderscore}}tac\ xs{\isaliteral{29}{\isacharparenright}}\isanewline |
17175 | 493 |
\isacommand{apply}\isamarkupfalse% |
40406 | 494 |
{\isaliteral{28}{\isacharparenleft}}auto{\isaliteral{29}{\isacharparenright}}\isanewline |
17175 | 495 |
\isacommand{done}\isamarkupfalse% |
496 |
% |
|
17056 | 497 |
\endisatagproof |
498 |
{\isafoldproof}% |
|
499 |
% |
|
500 |
\isadelimproof |
|
501 |
% |
|
502 |
\endisadelimproof |
|
11866 | 503 |
% |
8749 | 504 |
\begin{isamarkuptext}% |
505 |
\noindent |
|
11456 | 506 |
The final \commdx{end} tells Isabelle to close the current theory because |
8749 | 507 |
we are finished with its development:% |
11456 | 508 |
\index{*rev (constant)|)}\index{append function|)}% |
8749 | 509 |
\end{isamarkuptext}% |
17175 | 510 |
\isamarkuptrue% |
17056 | 511 |
% |
512 |
\isadelimtheory |
|
513 |
% |
|
514 |
\endisadelimtheory |
|
515 |
% |
|
516 |
\isatagtheory |
|
17175 | 517 |
\isacommand{end}\isamarkupfalse% |
518 |
% |
|
17056 | 519 |
\endisatagtheory |
520 |
{\isafoldtheory}% |
|
521 |
% |
|
522 |
\isadelimtheory |
|
523 |
% |
|
524 |
\endisadelimtheory |
|
525 |
\isanewline |
|
9722 | 526 |
\end{isabellebody}% |
9145 | 527 |
%%% Local Variables: |
528 |
%%% mode: latex |
|
529 |
%%% TeX-master: "root" |
|
530 |
%%% End: |