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