| 9924 |      1 | %
 | 
|  |      2 | \begin{isabellebody}%
 | 
|  |      3 | \def\isabellecontext{simp}%
 | 
| 17056 |      4 | %
 | 
|  |      5 | \isadelimtheory
 | 
|  |      6 | %
 | 
|  |      7 | \endisadelimtheory
 | 
|  |      8 | %
 | 
|  |      9 | \isatagtheory
 | 
|  |     10 | %
 | 
|  |     11 | \endisatagtheory
 | 
|  |     12 | {\isafoldtheory}%
 | 
|  |     13 | %
 | 
|  |     14 | \isadelimtheory
 | 
|  |     15 | %
 | 
|  |     16 | \endisadelimtheory
 | 
|  |     17 | \isamarkuptrue%
 | 
| 9933 |     18 | %
 | 
| 11214 |     19 | \isamarkupsubsection{Simplification Rules%
 | 
| 10397 |     20 | }
 | 
| 11866 |     21 | \isamarkuptrue%
 | 
| 9933 |     22 | %
 | 
|  |     23 | \begin{isamarkuptext}%
 | 
| 11458 |     24 | \index{simplification rules}
 | 
|  |     25 | To facilitate simplification,  
 | 
|  |     26 | the attribute \isa{{\isacharbrackleft}simp{\isacharbrackright}}\index{*simp (attribute)}
 | 
|  |     27 | declares theorems to be simplification rules, which the simplifier
 | 
|  |     28 | will use automatically.  In addition, \isacommand{datatype} and
 | 
|  |     29 | \isacommand{primrec} declarations (and a few others) 
 | 
|  |     30 | implicitly declare some simplification rules.  
 | 
|  |     31 | Explicit definitions are \emph{not} declared as 
 | 
| 9933 |     32 | simplification rules automatically!
 | 
|  |     33 | 
 | 
| 11458 |     34 | Nearly any theorem can become a simplification
 | 
|  |     35 | rule. The simplifier will try to transform it into an equation.  
 | 
|  |     36 | For example, the theorem
 | 
|  |     37 | \isa{{\isasymnot}\ P} is turned into \isa{P\ {\isacharequal}\ False}. The details
 | 
| 9933 |     38 | are explained in \S\ref{sec:SimpHow}.
 | 
|  |     39 | 
 | 
| 11458 |     40 | The simplification attribute of theorems can be turned on and off:%
 | 
| 12489 |     41 | \index{*simp del (attribute)}
 | 
| 9933 |     42 | \begin{quote}
 | 
|  |     43 | \isacommand{declare} \textit{theorem-name}\isa{{\isacharbrackleft}simp{\isacharbrackright}}\\
 | 
|  |     44 | \isacommand{declare} \textit{theorem-name}\isa{{\isacharbrackleft}simp\ del{\isacharbrackright}}
 | 
|  |     45 | \end{quote}
 | 
| 11309 |     46 | Only equations that really simplify, like \isa{rev\
 | 
|  |     47 | {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs} and
 | 
|  |     48 | \isa{xs\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}\
 | 
|  |     49 | {\isacharequal}\ xs}, should be declared as default simplification rules. 
 | 
|  |     50 | More specific ones should only be used selectively and should
 | 
|  |     51 | not be made default.  Distributivity laws, for example, alter
 | 
|  |     52 | the structure of terms and can produce an exponential blow-up instead of
 | 
|  |     53 | simplification.  A default simplification rule may
 | 
|  |     54 | need to be disabled in certain proofs.  Frequent changes in the simplification
 | 
|  |     55 | status of a theorem may indicate an unwise use of defaults.
 | 
| 9933 |     56 | \begin{warn}
 | 
| 11458 |     57 |   Simplification can run forever, for example if both $f(x) = g(x)$ and
 | 
| 9933 |     58 |   $g(x) = f(x)$ are simplification rules. It is the user's responsibility not
 | 
|  |     59 |   to include simplification rules that can lead to nontermination, either on
 | 
|  |     60 |   their own or in combination with other simplification rules.
 | 
| 12332 |     61 | \end{warn}
 | 
|  |     62 | \begin{warn}
 | 
|  |     63 |   It is inadvisable to toggle the simplification attribute of a
 | 
| 12333 |     64 |   theorem from a parent theory $A$ in a child theory $B$ for good.
 | 
| 12332 |     65 |   The reason is that if some theory $C$ is based both on $B$ and (via a
 | 
| 13814 |     66 |   different path) on $A$, it is not defined what the simplification attribute
 | 
| 12332 |     67 |   of that theorem will be in $C$: it could be either.
 | 
| 9933 |     68 | \end{warn}%
 | 
|  |     69 | \end{isamarkuptext}%
 | 
| 11866 |     70 | \isamarkuptrue%
 | 
| 9933 |     71 | %
 | 
| 11458 |     72 | \isamarkupsubsection{The {\tt\slshape simp}  Method%
 | 
| 10397 |     73 | }
 | 
| 11866 |     74 | \isamarkuptrue%
 | 
| 9933 |     75 | %
 | 
|  |     76 | \begin{isamarkuptext}%
 | 
|  |     77 | \index{*simp (method)|bold}
 | 
|  |     78 | The general format of the simplification method is
 | 
|  |     79 | \begin{quote}
 | 
|  |     80 | \isa{simp} \textit{list of modifiers}
 | 
|  |     81 | \end{quote}
 | 
| 10795 |     82 | where the list of \emph{modifiers} fine tunes the behaviour and may
 | 
| 11309 |     83 | be empty. Specific modifiers are discussed below.  Most if not all of the
 | 
|  |     84 | proofs seen so far could have been performed
 | 
| 9933 |     85 | with \isa{simp} instead of \isa{auto}, except that \isa{simp} attacks
 | 
| 10971 |     86 | only the first subgoal and may thus need to be repeated --- use
 | 
| 11428 |     87 | \methdx{simp_all} to simplify all subgoals.
 | 
| 11458 |     88 | If nothing changes, \isa{simp} fails.%
 | 
| 9933 |     89 | \end{isamarkuptext}%
 | 
| 11866 |     90 | \isamarkuptrue%
 | 
| 9933 |     91 | %
 | 
| 11214 |     92 | \isamarkupsubsection{Adding and Deleting Simplification Rules%
 | 
| 10397 |     93 | }
 | 
| 11866 |     94 | \isamarkuptrue%
 | 
| 9933 |     95 | %
 | 
|  |     96 | \begin{isamarkuptext}%
 | 
| 11458 |     97 | \index{simplification rules!adding and deleting}%
 | 
| 9933 |     98 | If a certain theorem is merely needed in a few proofs by simplification,
 | 
|  |     99 | we do not need to make it a global simplification rule. Instead we can modify
 | 
|  |    100 | the set of simplification rules used in a simplification step by adding rules
 | 
|  |    101 | to it and/or deleting rules from it. The two modifiers for this are
 | 
|  |    102 | \begin{quote}
 | 
| 11458 |    103 | \isa{add{\isacharcolon}} \textit{list of theorem names}\index{*add (modifier)}\\
 | 
|  |    104 | \isa{del{\isacharcolon}} \textit{list of theorem names}\index{*del (modifier)}
 | 
| 9933 |    105 | \end{quote}
 | 
| 11458 |    106 | Or you can use a specific list of theorems and omit all others:
 | 
| 9933 |    107 | \begin{quote}
 | 
| 11458 |    108 | \isa{only{\isacharcolon}} \textit{list of theorem names}\index{*only (modifier)}
 | 
| 11309 |    109 | \end{quote}
 | 
|  |    110 | In this example, we invoke the simplifier, adding two distributive
 | 
|  |    111 | laws:
 | 
|  |    112 | \begin{quote}
 | 
|  |    113 | \isacommand{apply}\isa{{\isacharparenleft}simp\ add{\isacharcolon}\ mod{\isacharunderscore}mult{\isacharunderscore}distrib\ add{\isacharunderscore}mult{\isacharunderscore}distrib{\isacharparenright}}
 | 
| 9933 |    114 | \end{quote}%
 | 
|  |    115 | \end{isamarkuptext}%
 | 
| 11866 |    116 | \isamarkuptrue%
 | 
| 9933 |    117 | %
 | 
| 11214 |    118 | \isamarkupsubsection{Assumptions%
 | 
| 10397 |    119 | }
 | 
| 11866 |    120 | \isamarkuptrue%
 | 
| 9933 |    121 | %
 | 
|  |    122 | \begin{isamarkuptext}%
 | 
|  |    123 | \index{simplification!with/of assumptions}
 | 
|  |    124 | By default, assumptions are part of the simplification process: they are used
 | 
|  |    125 | as simplification rules and are simplified themselves. For example:%
 | 
|  |    126 | \end{isamarkuptext}%
 | 
| 17056 |    127 | \isamarkupfalse%
 | 
| 9933 |    128 | \isacommand{lemma}\ {\isachardoublequote}{\isasymlbrakk}\ xs\ {\isacharat}\ zs\ {\isacharequal}\ ys\ {\isacharat}\ xs{\isacharsemicolon}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ ys\ {\isacharequal}\ zs{\isachardoublequote}\isanewline
 | 
| 17056 |    129 | %
 | 
|  |    130 | \isadelimproof
 | 
|  |    131 | %
 | 
|  |    132 | \endisadelimproof
 | 
|  |    133 | %
 | 
|  |    134 | \isatagproof
 | 
| 11866 |    135 | \isamarkupfalse%
 | 
| 16069 |    136 | \isacommand{apply}\ simp\isanewline
 | 
| 11866 |    137 | \isamarkupfalse%
 | 
| 17056 |    138 | \isacommand{done}%
 | 
|  |    139 | \endisatagproof
 | 
|  |    140 | {\isafoldproof}%
 | 
|  |    141 | %
 | 
|  |    142 | \isadelimproof
 | 
|  |    143 | %
 | 
|  |    144 | \endisadelimproof
 | 
|  |    145 | \isamarkuptrue%
 | 
| 11866 |    146 | %
 | 
| 9933 |    147 | \begin{isamarkuptext}%
 | 
|  |    148 | \noindent
 | 
|  |    149 | The second assumption simplifies to \isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}}, which in turn
 | 
|  |    150 | simplifies the first assumption to \isa{zs\ {\isacharequal}\ ys}, thus reducing the
 | 
|  |    151 | conclusion to \isa{ys\ {\isacharequal}\ ys} and hence to \isa{True}.
 | 
|  |    152 | 
 | 
| 11458 |    153 | In some cases, using the assumptions can lead to nontermination:%
 | 
| 9933 |    154 | \end{isamarkuptext}%
 | 
| 17056 |    155 | \isamarkupfalse%
 | 
|  |    156 | \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x{\isachardot}\ f\ x\ {\isacharequal}\ g\ {\isacharparenleft}f\ {\isacharparenleft}g\ x{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ f\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}%
 | 
|  |    157 | \isadelimproof
 | 
|  |    158 | %
 | 
|  |    159 | \endisadelimproof
 | 
|  |    160 | %
 | 
|  |    161 | \isatagproof
 | 
| 11866 |    162 | \isamarkuptrue%
 | 
| 16069 |    163 | %
 | 
|  |    164 | \begin{isamarkuptxt}%
 | 
|  |    165 | \noindent
 | 
|  |    166 | An unmodified application of \isa{simp} loops.  The culprit is the
 | 
|  |    167 | simplification rule \isa{f\ x\ {\isacharequal}\ g\ {\isacharparenleft}f\ {\isacharparenleft}g\ x{\isacharparenright}{\isacharparenright}}, which is extracted from
 | 
|  |    168 | the assumption.  (Isabelle notices certain simple forms of
 | 
|  |    169 | nontermination but not this one.)  The problem can be circumvented by
 | 
|  |    170 | telling the simplifier to ignore the assumptions:%
 | 
|  |    171 | \end{isamarkuptxt}%
 | 
| 17056 |    172 | \isamarkupfalse%
 | 
| 16069 |    173 | \isacommand{apply}{\isacharparenleft}simp\ {\isacharparenleft}no{\isacharunderscore}asm{\isacharparenright}{\isacharparenright}\isanewline
 | 
| 11866 |    174 | \isamarkupfalse%
 | 
| 17056 |    175 | \isacommand{done}%
 | 
|  |    176 | \endisatagproof
 | 
|  |    177 | {\isafoldproof}%
 | 
|  |    178 | %
 | 
|  |    179 | \isadelimproof
 | 
|  |    180 | %
 | 
|  |    181 | \endisadelimproof
 | 
|  |    182 | \isamarkuptrue%
 | 
| 11866 |    183 | %
 | 
| 9933 |    184 | \begin{isamarkuptext}%
 | 
|  |    185 | \noindent
 | 
| 11458 |    186 | Three modifiers influence the treatment of assumptions:
 | 
| 9933 |    187 | \begin{description}
 | 
| 11458 |    188 | \item[\isa{{\isacharparenleft}no{\isacharunderscore}asm{\isacharparenright}}]\index{*no_asm (modifier)}
 | 
| 9933 |    189 |  means that assumptions are completely ignored.
 | 
| 11458 |    190 | \item[\isa{{\isacharparenleft}no{\isacharunderscore}asm{\isacharunderscore}simp{\isacharparenright}}]\index{*no_asm_simp (modifier)}
 | 
| 9933 |    191 |  means that the assumptions are not simplified but
 | 
|  |    192 |   are used in the simplification of the conclusion.
 | 
| 11458 |    193 | \item[\isa{{\isacharparenleft}no{\isacharunderscore}asm{\isacharunderscore}use{\isacharparenright}}]\index{*no_asm_use (modifier)}
 | 
| 9933 |    194 |  means that the assumptions are simplified but are not
 | 
|  |    195 |   used in the simplification of each other or the conclusion.
 | 
|  |    196 | \end{description}
 | 
| 10795 |    197 | Both \isa{{\isacharparenleft}no{\isacharunderscore}asm{\isacharunderscore}simp{\isacharparenright}} and \isa{{\isacharparenleft}no{\isacharunderscore}asm{\isacharunderscore}use{\isacharparenright}} run forever on
 | 
|  |    198 | the problematic subgoal above.
 | 
| 11458 |    199 | Only one of the modifiers is allowed, and it must precede all
 | 
| 11309 |    200 | other modifiers.
 | 
| 13623 |    201 | %\begin{warn}
 | 
|  |    202 | %Assumptions are simplified in a left-to-right fashion. If an
 | 
|  |    203 | %assumption can help in simplifying one to the left of it, this may get
 | 
|  |    204 | %overlooked. In such cases you have to rotate the assumptions explicitly:
 | 
|  |    205 | %\isacommand{apply}@ {text"("}\methdx{rotate_tac}~$n$@ {text")"}
 | 
|  |    206 | %causes a cyclic shift by $n$ positions from right to left, if $n$ is
 | 
|  |    207 | %positive, and from left to right, if $n$ is negative.
 | 
|  |    208 | %Beware that such rotations make proofs quite brittle.
 | 
|  |    209 | %\end{warn}%
 | 
| 9933 |    210 | \end{isamarkuptext}%
 | 
| 11866 |    211 | \isamarkuptrue%
 | 
| 9933 |    212 | %
 | 
| 11214 |    213 | \isamarkupsubsection{Rewriting with Definitions%
 | 
| 10397 |    214 | }
 | 
| 11866 |    215 | \isamarkuptrue%
 | 
| 9933 |    216 | %
 | 
|  |    217 | \begin{isamarkuptext}%
 | 
| 11215 |    218 | \label{sec:Simp-with-Defs}\index{simplification!with definitions}
 | 
| 11458 |    219 | Constant definitions (\S\ref{sec:ConstDefinitions}) can be used as
 | 
|  |    220 | simplification rules, but by default they are not: the simplifier does not
 | 
|  |    221 | expand them automatically.  Definitions are intended for introducing abstract
 | 
| 12584 |    222 | concepts and not merely as abbreviations.  Of course, we need to expand
 | 
| 11458 |    223 | the definition initially, but once we have proved enough abstract properties
 | 
|  |    224 | of the new constant, we can forget its original definition.  This style makes
 | 
|  |    225 | proofs more robust: if the definition has to be changed,
 | 
|  |    226 | only the proofs of the abstract properties will be affected.
 | 
|  |    227 | 
 | 
|  |    228 | For example, given%
 | 
| 9933 |    229 | \end{isamarkuptext}%
 | 
| 17056 |    230 | \isamarkupfalse%
 | 
| 10788 |    231 | \isacommand{constdefs}\ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\isanewline
 | 
| 17056 |    232 | \ \ \ \ \ \ \ \ \ {\isachardoublequote}xor\ A\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isamarkuptrue%
 | 
| 11866 |    233 | %
 | 
| 9933 |    234 | \begin{isamarkuptext}%
 | 
|  |    235 | \noindent
 | 
|  |    236 | we may want to prove%
 | 
|  |    237 | \end{isamarkuptext}%
 | 
| 17056 |    238 | \isamarkupfalse%
 | 
|  |    239 | \isacommand{lemma}\ {\isachardoublequote}xor\ A\ {\isacharparenleft}{\isasymnot}A{\isacharparenright}{\isachardoublequote}%
 | 
|  |    240 | \isadelimproof
 | 
|  |    241 | %
 | 
|  |    242 | \endisadelimproof
 | 
|  |    243 | %
 | 
|  |    244 | \isatagproof
 | 
| 11866 |    245 | \isamarkuptrue%
 | 
| 16069 |    246 | %
 | 
|  |    247 | \begin{isamarkuptxt}%
 | 
|  |    248 | \noindent
 | 
|  |    249 | Typically, we begin by unfolding some definitions:
 | 
|  |    250 | \indexbold{definitions!unfolding}%
 | 
|  |    251 | \end{isamarkuptxt}%
 | 
| 17056 |    252 | \isamarkupfalse%
 | 
|  |    253 | \isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ xor{\isacharunderscore}def{\isacharparenright}\isamarkuptrue%
 | 
| 16069 |    254 | %
 | 
|  |    255 | \begin{isamarkuptxt}%
 | 
|  |    256 | \noindent
 | 
|  |    257 | In this particular case, the resulting goal
 | 
|  |    258 | \begin{isabelle}%
 | 
|  |    259 | \ {\isadigit{1}}{\isachardot}\ A\ {\isasymand}\ {\isasymnot}\ {\isasymnot}\ A\ {\isasymor}\ {\isasymnot}\ A\ {\isasymand}\ {\isasymnot}\ A%
 | 
|  |    260 | \end{isabelle}
 | 
|  |    261 | can be proved by simplification. Thus we could have proved the lemma outright by%
 | 
|  |    262 | \end{isamarkuptxt}%
 | 
| 17056 |    263 | %
 | 
|  |    264 | \endisatagproof
 | 
|  |    265 | {\isafoldproof}%
 | 
|  |    266 | %
 | 
|  |    267 | \isadelimproof
 | 
|  |    268 | %
 | 
|  |    269 | \endisadelimproof
 | 
|  |    270 | %
 | 
|  |    271 | \isadelimproof
 | 
|  |    272 | %
 | 
|  |    273 | \endisadelimproof
 | 
|  |    274 | %
 | 
|  |    275 | \isatagproof
 | 
| 11866 |    276 | \isamarkupfalse%
 | 
| 17056 |    277 | \isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ xor{\isacharunderscore}def{\isacharparenright}%
 | 
|  |    278 | \endisatagproof
 | 
|  |    279 | {\isafoldproof}%
 | 
|  |    280 | %
 | 
|  |    281 | \isadelimproof
 | 
|  |    282 | %
 | 
|  |    283 | \endisadelimproof
 | 
|  |    284 | \isamarkuptrue%
 | 
| 11866 |    285 | %
 | 
| 9933 |    286 | \begin{isamarkuptext}%
 | 
|  |    287 | \noindent
 | 
|  |    288 | Of course we can also unfold definitions in the middle of a proof.
 | 
|  |    289 | 
 | 
|  |    290 | \begin{warn}
 | 
| 10971 |    291 |   If you have defined $f\,x\,y~\isasymequiv~t$ then you can only unfold
 | 
|  |    292 |   occurrences of $f$ with at least two arguments. This may be helpful for unfolding
 | 
|  |    293 |   $f$ selectively, but it may also get in the way. Defining
 | 
|  |    294 |   $f$~\isasymequiv~\isasymlambda$x\,y.\;t$ allows to unfold all occurrences of $f$.
 | 
| 12473 |    295 | \end{warn}
 | 
|  |    296 | 
 | 
|  |    297 | There is also the special method \isa{unfold}\index{*unfold (method)|bold}
 | 
|  |    298 | which merely unfolds
 | 
|  |    299 | one or several definitions, as in \isacommand{apply}\isa{(unfold xor_def)}.
 | 
| 12533 |    300 | This is can be useful in situations where \isa{simp} does too much.
 | 
|  |    301 | Warning: \isa{unfold} acts on all subgoals!%
 | 
| 9933 |    302 | \end{isamarkuptext}%
 | 
| 11866 |    303 | \isamarkuptrue%
 | 
| 9933 |    304 | %
 | 
| 11214 |    305 | \isamarkupsubsection{Simplifying {\tt\slshape let}-Expressions%
 | 
| 10397 |    306 | }
 | 
| 11866 |    307 | \isamarkuptrue%
 | 
| 9933 |    308 | %
 | 
|  |    309 | \begin{isamarkuptext}%
 | 
| 11458 |    310 | \index{simplification!of \isa{let}-expressions}\index{*let expressions}%
 | 
|  |    311 | Proving a goal containing \isa{let}-expressions almost invariably requires the
 | 
|  |    312 | \isa{let}-con\-structs to be expanded at some point. Since
 | 
|  |    313 | \isa{let}\ldots\isa{=}\ldots\isa{in}{\ldots} is just syntactic sugar for
 | 
|  |    314 | the predefined constant \isa{Let}, expanding \isa{let}-constructs
 | 
|  |    315 | means rewriting with \tdx{Let_def}:%
 | 
| 9933 |    316 | \end{isamarkuptext}%
 | 
| 17056 |    317 | \isamarkupfalse%
 | 
| 9933 |    318 | \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}let\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ in\ xs{\isacharat}ys{\isacharat}xs{\isacharparenright}\ {\isacharequal}\ ys{\isachardoublequote}\isanewline
 | 
| 17056 |    319 | %
 | 
|  |    320 | \isadelimproof
 | 
|  |    321 | %
 | 
|  |    322 | \endisadelimproof
 | 
|  |    323 | %
 | 
|  |    324 | \isatagproof
 | 
| 11866 |    325 | \isamarkupfalse%
 | 
| 16069 |    326 | \isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ Let{\isacharunderscore}def{\isacharparenright}\isanewline
 | 
| 11866 |    327 | \isamarkupfalse%
 | 
| 17056 |    328 | \isacommand{done}%
 | 
|  |    329 | \endisatagproof
 | 
|  |    330 | {\isafoldproof}%
 | 
|  |    331 | %
 | 
|  |    332 | \isadelimproof
 | 
|  |    333 | %
 | 
|  |    334 | \endisadelimproof
 | 
|  |    335 | \isamarkuptrue%
 | 
| 11866 |    336 | %
 | 
| 9933 |    337 | \begin{isamarkuptext}%
 | 
|  |    338 | If, in a particular context, there is no danger of a combinatorial explosion
 | 
| 11458 |    339 | of nested \isa{let}s, you could even simplify with \isa{Let{\isacharunderscore}def} by
 | 
| 9933 |    340 | default:%
 | 
|  |    341 | \end{isamarkuptext}%
 | 
| 17056 |    342 | \isamarkupfalse%
 | 
|  |    343 | \isacommand{declare}\ Let{\isacharunderscore}def\ {\isacharbrackleft}simp{\isacharbrackright}\isamarkuptrue%
 | 
| 11866 |    344 | %
 | 
| 11458 |    345 | \isamarkupsubsection{Conditional Simplification Rules%
 | 
| 10397 |    346 | }
 | 
| 11866 |    347 | \isamarkuptrue%
 | 
| 9933 |    348 | %
 | 
|  |    349 | \begin{isamarkuptext}%
 | 
| 11458 |    350 | \index{conditional simplification rules}%
 | 
| 9933 |    351 | So far all examples of rewrite rules were equations. The simplifier also
 | 
|  |    352 | accepts \emph{conditional} equations, for example%
 | 
|  |    353 | \end{isamarkuptext}%
 | 
| 17056 |    354 | \isamarkupfalse%
 | 
| 9933 |    355 | \isacommand{lemma}\ hd{\isacharunderscore}Cons{\isacharunderscore}tl{\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ \ {\isasymLongrightarrow}\ \ hd\ xs\ {\isacharhash}\ tl\ xs\ {\isacharequal}\ xs{\isachardoublequote}\isanewline
 | 
| 17056 |    356 | %
 | 
|  |    357 | \isadelimproof
 | 
|  |    358 | %
 | 
|  |    359 | \endisadelimproof
 | 
|  |    360 | %
 | 
|  |    361 | \isatagproof
 | 
| 11866 |    362 | \isamarkupfalse%
 | 
| 16069 |    363 | \isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharcomma}\ simp{\isacharparenright}\isanewline
 | 
| 11866 |    364 | \isamarkupfalse%
 | 
| 17056 |    365 | \isacommand{done}%
 | 
|  |    366 | \endisatagproof
 | 
|  |    367 | {\isafoldproof}%
 | 
|  |    368 | %
 | 
|  |    369 | \isadelimproof
 | 
|  |    370 | %
 | 
|  |    371 | \endisadelimproof
 | 
|  |    372 | \isamarkuptrue%
 | 
| 11866 |    373 | %
 | 
| 9933 |    374 | \begin{isamarkuptext}%
 | 
|  |    375 | \noindent
 | 
|  |    376 | Note the use of ``\ttindexboldpos{,}{$Isar}'' to string together a
 | 
|  |    377 | sequence of methods. Assuming that the simplification rule
 | 
|  |    378 | \isa{{\isacharparenleft}rev\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}}
 | 
| 10795 |    379 | is present as well,
 | 
|  |    380 | the lemma below is proved by plain simplification:%
 | 
| 9933 |    381 | \end{isamarkuptext}%
 | 
| 17056 |    382 | \isamarkupfalse%
 | 
|  |    383 | \isacommand{lemma}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharhash}\ tl{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ rev\ xs{\isachardoublequote}%
 | 
|  |    384 | \isadelimproof
 | 
|  |    385 | %
 | 
|  |    386 | \endisadelimproof
 | 
|  |    387 | %
 | 
|  |    388 | \isatagproof
 | 
|  |    389 | %
 | 
|  |    390 | \endisatagproof
 | 
|  |    391 | {\isafoldproof}%
 | 
|  |    392 | %
 | 
|  |    393 | \isadelimproof
 | 
|  |    394 | %
 | 
|  |    395 | \endisadelimproof
 | 
| 11866 |    396 | \isamarkuptrue%
 | 
|  |    397 | %
 | 
| 9933 |    398 | \begin{isamarkuptext}%
 | 
|  |    399 | \noindent
 | 
| 10795 |    400 | The conditional equation \isa{hd{\isacharunderscore}Cons{\isacharunderscore}tl} above
 | 
| 9933 |    401 | can simplify \isa{hd\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharhash}\ tl\ {\isacharparenleft}rev\ xs{\isacharparenright}} to \isa{rev\ xs}
 | 
|  |    402 | because the corresponding precondition \isa{rev\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}}
 | 
|  |    403 | simplifies to \isa{xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}}, which is exactly the local
 | 
|  |    404 | assumption of the subgoal.%
 | 
|  |    405 | \end{isamarkuptext}%
 | 
| 11866 |    406 | \isamarkuptrue%
 | 
| 9933 |    407 | %
 | 
| 11214 |    408 | \isamarkupsubsection{Automatic Case Splits%
 | 
| 10397 |    409 | }
 | 
| 11866 |    410 | \isamarkuptrue%
 | 
| 9933 |    411 | %
 | 
|  |    412 | \begin{isamarkuptext}%
 | 
| 11458 |    413 | \label{sec:AutoCaseSplits}\indexbold{case splits}%
 | 
|  |    414 | Goals containing \isa{if}-expressions\index{*if expressions!splitting of}
 | 
|  |    415 | are usually proved by case
 | 
|  |    416 | distinction on the boolean condition.  Here is an example:%
 | 
| 9933 |    417 | \end{isamarkuptext}%
 | 
| 17056 |    418 | \isamarkupfalse%
 | 
|  |    419 | \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}xs{\isachardot}\ if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ rev\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ else\ rev\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}%
 | 
|  |    420 | \isadelimproof
 | 
|  |    421 | %
 | 
|  |    422 | \endisadelimproof
 | 
|  |    423 | %
 | 
|  |    424 | \isatagproof
 | 
| 11866 |    425 | \isamarkuptrue%
 | 
| 16069 |    426 | %
 | 
|  |    427 | \begin{isamarkuptxt}%
 | 
|  |    428 | \noindent
 | 
|  |    429 | The goal can be split by a special method, \methdx{split}:%
 | 
|  |    430 | \end{isamarkuptxt}%
 | 
| 17056 |    431 | \isamarkupfalse%
 | 
|  |    432 | \isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}if{\isacharparenright}\isamarkuptrue%
 | 
| 16069 |    433 | %
 | 
|  |    434 | \begin{isamarkuptxt}%
 | 
|  |    435 | \noindent
 | 
|  |    436 | \begin{isabelle}%
 | 
|  |    437 | \ {\isadigit{1}}{\isachardot}\ {\isasymforall}xs{\isachardot}\ {\isacharparenleft}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ rev\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ rev\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}%
 | 
|  |    438 | \end{isabelle}
 | 
|  |    439 | where \tdx{split_if} is a theorem that expresses splitting of
 | 
|  |    440 | \isa{if}s. Because
 | 
|  |    441 | splitting the \isa{if}s is usually the right proof strategy, the
 | 
|  |    442 | simplifier does it automatically.  Try \isacommand{apply}\isa{{\isacharparenleft}simp{\isacharparenright}}
 | 
|  |    443 | on the initial goal above.
 | 
|  |    444 | 
 | 
|  |    445 | This splitting idea generalizes from \isa{if} to \sdx{case}.
 | 
|  |    446 | Let us simplify a case analysis over lists:\index{*list.split (theorem)}%
 | 
|  |    447 | \end{isamarkuptxt}%
 | 
| 17056 |    448 | %
 | 
|  |    449 | \endisatagproof
 | 
|  |    450 | {\isafoldproof}%
 | 
|  |    451 | %
 | 
|  |    452 | \isadelimproof
 | 
|  |    453 | %
 | 
|  |    454 | \endisadelimproof
 | 
| 11866 |    455 | \isamarkupfalse%
 | 
| 10362 |    456 | \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ zs\ {\isacharbar}\ y{\isacharhash}ys\ {\isasymRightarrow}\ y{\isacharhash}{\isacharparenleft}ys{\isacharat}zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ xs{\isacharat}zs{\isachardoublequote}\isanewline
 | 
| 17056 |    457 | %
 | 
|  |    458 | \isadelimproof
 | 
|  |    459 | %
 | 
|  |    460 | \endisadelimproof
 | 
|  |    461 | %
 | 
|  |    462 | \isatagproof
 | 
| 11866 |    463 | \isamarkupfalse%
 | 
| 17056 |    464 | \isacommand{apply}{\isacharparenleft}split\ list{\isachardot}split{\isacharparenright}\isamarkuptrue%
 | 
| 16069 |    465 | %
 | 
|  |    466 | \begin{isamarkuptxt}%
 | 
|  |    467 | \begin{isabelle}%
 | 
|  |    468 | \ {\isadigit{1}}{\isachardot}\ {\isacharparenleft}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ zs\ {\isacharequal}\ xs\ {\isacharat}\ zs{\isacharparenright}\ {\isasymand}\isanewline
 | 
|  |    469 | \isaindent{\ {\isadigit{1}}{\isachardot}\ }{\isacharparenleft}{\isasymforall}a\ list{\isachardot}\ xs\ {\isacharequal}\ a\ {\isacharhash}\ list\ {\isasymlongrightarrow}\ a\ {\isacharhash}\ list\ {\isacharat}\ zs\ {\isacharequal}\ xs\ {\isacharat}\ zs{\isacharparenright}%
 | 
|  |    470 | \end{isabelle}
 | 
|  |    471 | The simplifier does not split
 | 
|  |    472 | \isa{case}-expressions, as it does \isa{if}-expressions, 
 | 
|  |    473 | because with recursive datatypes it could lead to nontermination.
 | 
|  |    474 | Instead, the simplifier has a modifier
 | 
|  |    475 | \isa{split}\index{*split (modifier)} 
 | 
|  |    476 | for adding splitting rules explicitly.  The
 | 
|  |    477 | lemma above can be proved in one step by%
 | 
|  |    478 | \end{isamarkuptxt}%
 | 
| 17056 |    479 | %
 | 
|  |    480 | \endisatagproof
 | 
|  |    481 | {\isafoldproof}%
 | 
|  |    482 | %
 | 
|  |    483 | \isadelimproof
 | 
|  |    484 | %
 | 
|  |    485 | \endisadelimproof
 | 
|  |    486 | %
 | 
|  |    487 | \isadelimproof
 | 
|  |    488 | %
 | 
|  |    489 | \endisadelimproof
 | 
|  |    490 | %
 | 
|  |    491 | \isatagproof
 | 
| 11866 |    492 | \isamarkupfalse%
 | 
| 17056 |    493 | \isacommand{apply}{\isacharparenleft}simp\ split{\isacharcolon}\ list{\isachardot}split{\isacharparenright}%
 | 
|  |    494 | \endisatagproof
 | 
|  |    495 | {\isafoldproof}%
 | 
|  |    496 | %
 | 
|  |    497 | \isadelimproof
 | 
|  |    498 | %
 | 
|  |    499 | \endisadelimproof
 | 
|  |    500 | \isamarkuptrue%
 | 
| 11866 |    501 | %
 | 
| 9933 |    502 | \begin{isamarkuptext}%
 | 
| 10668 |    503 | \noindent
 | 
|  |    504 | whereas \isacommand{apply}\isa{{\isacharparenleft}simp{\isacharparenright}} alone will not succeed.
 | 
| 9933 |    505 | 
 | 
| 11458 |    506 | Every datatype $t$ comes with a theorem
 | 
| 9933 |    507 | $t$\isa{{\isachardot}split} which can be declared to be a \bfindex{split rule} either
 | 
| 11458 |    508 | locally as above, or by giving it the \attrdx{split} attribute globally:%
 | 
| 9933 |    509 | \end{isamarkuptext}%
 | 
| 17056 |    510 | \isamarkupfalse%
 | 
|  |    511 | \isacommand{declare}\ list{\isachardot}split\ {\isacharbrackleft}split{\isacharbrackright}\isamarkuptrue%
 | 
| 11866 |    512 | %
 | 
| 9933 |    513 | \begin{isamarkuptext}%
 | 
|  |    514 | \noindent
 | 
|  |    515 | The \isa{split} attribute can be removed with the \isa{del} modifier,
 | 
|  |    516 | either locally%
 | 
|  |    517 | \end{isamarkuptext}%
 | 
| 17056 |    518 | %
 | 
|  |    519 | \isadelimproof
 | 
|  |    520 | %
 | 
|  |    521 | \endisadelimproof
 | 
|  |    522 | %
 | 
|  |    523 | \isatagproof
 | 
| 11866 |    524 | \isamarkupfalse%
 | 
| 17056 |    525 | \isacommand{apply}{\isacharparenleft}simp\ split\ del{\isacharcolon}\ split{\isacharunderscore}if{\isacharparenright}%
 | 
|  |    526 | \endisatagproof
 | 
|  |    527 | {\isafoldproof}%
 | 
|  |    528 | %
 | 
|  |    529 | \isadelimproof
 | 
|  |    530 | %
 | 
|  |    531 | \endisadelimproof
 | 
|  |    532 | \isamarkuptrue%
 | 
| 11866 |    533 | %
 | 
| 9933 |    534 | \begin{isamarkuptext}%
 | 
|  |    535 | \noindent
 | 
|  |    536 | or globally:%
 | 
|  |    537 | \end{isamarkuptext}%
 | 
| 17056 |    538 | \isamarkupfalse%
 | 
|  |    539 | \isacommand{declare}\ list{\isachardot}split\ {\isacharbrackleft}split\ del{\isacharbrackright}\isamarkuptrue%
 | 
| 11866 |    540 | %
 | 
| 9933 |    541 | \begin{isamarkuptext}%
 | 
| 11458 |    542 | Polished proofs typically perform splitting within \isa{simp} rather than 
 | 
|  |    543 | invoking the \isa{split} method.  However, if a goal contains
 | 
|  |    544 | several \isa{if} and \isa{case} expressions, 
 | 
|  |    545 | the \isa{split} method can be
 | 
| 10668 |    546 | helpful in selectively exploring the effects of splitting.
 | 
|  |    547 | 
 | 
| 11458 |    548 | The split rules shown above are intended to affect only the subgoal's
 | 
|  |    549 | conclusion.  If you want to split an \isa{if} or \isa{case}-expression
 | 
|  |    550 | in the assumptions, you have to apply \tdx{split_if_asm} or
 | 
| 9933 |    551 | $t$\isa{{\isachardot}split{\isacharunderscore}asm}:%
 | 
|  |    552 | \end{isamarkuptext}%
 | 
| 17056 |    553 | \isamarkupfalse%
 | 
| 10668 |    554 | \isacommand{lemma}\ {\isachardoublequote}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ else\ ys\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ xs\ {\isacharat}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline
 | 
| 17056 |    555 | %
 | 
|  |    556 | \isadelimproof
 | 
|  |    557 | %
 | 
|  |    558 | \endisadelimproof
 | 
|  |    559 | %
 | 
|  |    560 | \isatagproof
 | 
| 11866 |    561 | \isamarkupfalse%
 | 
| 17056 |    562 | \isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}if{\isacharunderscore}asm{\isacharparenright}\isamarkuptrue%
 | 
| 16069 |    563 | %
 | 
|  |    564 | \begin{isamarkuptxt}%
 | 
|  |    565 | \noindent
 | 
|  |    566 | Unlike splitting the conclusion, this step creates two
 | 
|  |    567 | separate subgoals, which here can be solved by \isa{simp{\isacharunderscore}all}:
 | 
|  |    568 | \begin{isabelle}%
 | 
|  |    569 | \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ xs\ {\isacharat}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\isanewline
 | 
|  |    570 | \ {\isadigit{2}}{\isachardot}\ {\isasymlbrakk}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\ ys\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ xs\ {\isacharat}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}%
 | 
|  |    571 | \end{isabelle}
 | 
|  |    572 | If you need to split both in the assumptions and the conclusion,
 | 
|  |    573 | use $t$\isa{{\isachardot}splits} which subsumes $t$\isa{{\isachardot}split} and
 | 
|  |    574 | $t$\isa{{\isachardot}split{\isacharunderscore}asm}. Analogously, there is \isa{if{\isacharunderscore}splits}.
 | 
|  |    575 | 
 | 
|  |    576 | \begin{warn}
 | 
|  |    577 |   The simplifier merely simplifies the condition of an 
 | 
|  |    578 |   \isa{if}\index{*if expressions!simplification of} but not the
 | 
|  |    579 |   \isa{then} or \isa{else} parts. The latter are simplified only after the
 | 
|  |    580 |   condition reduces to \isa{True} or \isa{False}, or after splitting. The
 | 
|  |    581 |   same is true for \sdx{case}-expressions: only the selector is
 | 
|  |    582 |   simplified at first, until either the expression reduces to one of the
 | 
|  |    583 |   cases or it is split.
 | 
|  |    584 | \end{warn}%
 | 
|  |    585 | \end{isamarkuptxt}%
 | 
| 17056 |    586 | %
 | 
|  |    587 | \endisatagproof
 | 
|  |    588 | {\isafoldproof}%
 | 
|  |    589 | %
 | 
|  |    590 | \isadelimproof
 | 
|  |    591 | %
 | 
|  |    592 | \endisadelimproof
 | 
| 11866 |    593 | \isamarkuptrue%
 | 
| 9933 |    594 | %
 | 
| 11214 |    595 | \isamarkupsubsection{Tracing%
 | 
| 10397 |    596 | }
 | 
| 11866 |    597 | \isamarkuptrue%
 | 
| 9933 |    598 | %
 | 
|  |    599 | \begin{isamarkuptext}%
 | 
|  |    600 | \indexbold{tracing the simplifier}
 | 
|  |    601 | Using the simplifier effectively may take a bit of experimentation.  Set the
 | 
| 16523 |    602 | Proof General flag \pgmenu{Isabelle} $>$ \pgmenu{Settings} $>$ \pgmenu{Trace Simplifier} to get a better idea of what is going on:%
 | 
| 9933 |    603 | \end{isamarkuptext}%
 | 
| 17056 |    604 | \isamarkupfalse%
 | 
| 9933 |    605 | \isacommand{lemma}\ {\isachardoublequote}rev\ {\isacharbrackleft}a{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline
 | 
| 17056 |    606 | %
 | 
|  |    607 | \isadelimproof
 | 
|  |    608 | %
 | 
|  |    609 | \endisadelimproof
 | 
|  |    610 | %
 | 
|  |    611 | \isatagproof
 | 
| 11866 |    612 | \isamarkupfalse%
 | 
| 17056 |    613 | \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}%
 | 
|  |    614 | \endisatagproof
 | 
|  |    615 | {\isafoldproof}%
 | 
|  |    616 | %
 | 
|  |    617 | \isadelimproof
 | 
|  |    618 | %
 | 
|  |    619 | \endisadelimproof
 | 
|  |    620 | \isamarkuptrue%
 | 
| 11866 |    621 | %
 | 
| 9933 |    622 | \begin{isamarkuptext}%
 | 
|  |    623 | \noindent
 | 
| 16523 |    624 | produces the following trace in Proof General's \pgmenu{Trace} buffer:
 | 
| 9933 |    625 | 
 | 
|  |    626 | \begin{ttbox}\makeatother
 | 
| 16519 |    627 | [1]Applying instance of rewrite rule "List.rev.simps_2":
 | 
| 16359 |    628 | rev (?x1 # ?xs1) \(\equiv\) rev ?xs1 @ [?x1]
 | 
|  |    629 | 
 | 
| 16519 |    630 | [1]Rewriting:
 | 
| 16359 |    631 | rev [a] \(\equiv\) rev [] @ [a]
 | 
|  |    632 | 
 | 
| 16519 |    633 | [1]Applying instance of rewrite rule "List.rev.simps_1":
 | 
| 16359 |    634 | rev [] \(\equiv\) []
 | 
|  |    635 | 
 | 
| 16519 |    636 | [1]Rewriting:
 | 
| 16359 |    637 | rev [] \(\equiv\) []
 | 
|  |    638 | 
 | 
| 16519 |    639 | [1]Applying instance of rewrite rule "List.op @.append_Nil":
 | 
| 16359 |    640 | [] @ ?y \(\equiv\) ?y
 | 
|  |    641 | 
 | 
| 16519 |    642 | [1]Rewriting:
 | 
| 16359 |    643 | [] @ [a] \(\equiv\) [a]
 | 
|  |    644 | 
 | 
| 16519 |    645 | [1]Applying instance of rewrite rule
 | 
| 16359 |    646 | ?x2 # ?t1 = ?t1 \(\equiv\) False
 | 
|  |    647 | 
 | 
| 16519 |    648 | [1]Rewriting:
 | 
| 16359 |    649 | [a] = [] \(\equiv\) False
 | 
| 9933 |    650 | \end{ttbox}
 | 
| 16359 |    651 | The trace lists each rule being applied, both in its general form and
 | 
|  |    652 | the instance being used. The \texttt{[}$i$\texttt{]} in front (where
 | 
| 16519 |    653 | above $i$ is always \texttt{1}) indicates that we are inside the $i$th
 | 
|  |    654 | invocation of the simplifier. Each attempt to apply a
 | 
| 16359 |    655 | conditional rule shows the rule followed by the trace of the
 | 
|  |    656 | (recursive!) simplification of the conditions, the latter prefixed by
 | 
|  |    657 | \texttt{[}$i+1$\texttt{]} instead of \texttt{[}$i$\texttt{]}.
 | 
|  |    658 | Another source of recursive invocations of the simplifier are
 | 
|  |    659 | proofs of arithmetic formulae.
 | 
| 9933 |    660 | 
 | 
| 16359 |    661 | Many other hints about the simplifier's actions may appear.
 | 
| 11309 |    662 | 
 | 
| 16359 |    663 | In more complicated cases, the trace can be very lengthy.  Thus it is
 | 
| 16523 |    664 | advisable to reset the \pgmenu{Trace Simplifier} flag after having
 | 
| 16359 |    665 | obtained the desired trace.%
 | 
| 9933 |    666 | \end{isamarkuptext}%
 | 
| 11866 |    667 | \isamarkuptrue%
 | 
| 16519 |    668 | %
 | 
| 16545 |    669 | \isamarkupsubsection{Finding Theorems\label{sec:find}%
 | 
| 16519 |    670 | }
 | 
|  |    671 | \isamarkuptrue%
 | 
|  |    672 | %
 | 
|  |    673 | \begin{isamarkuptext}%
 | 
| 16523 |    674 | \indexbold{finding theorems}\indexbold{searching theorems}
 | 
| 16560 |    675 | Isabelle's large database of proved theorems 
 | 
|  |    676 | offers a powerful search engine. Its chief limitation is
 | 
| 16519 |    677 | its restriction to the theories currently loaded.
 | 
|  |    678 | 
 | 
|  |    679 | \begin{pgnote}
 | 
| 16523 |    680 | The search engine is started by clicking on Proof General's \pgmenu{Find} icon.
 | 
| 16519 |    681 | You specify your search textually in the input buffer at the bottom
 | 
|  |    682 | of the window.
 | 
|  |    683 | \end{pgnote}
 | 
|  |    684 | 
 | 
| 16560 |    685 | The simplest form of search finds theorems containing specified
 | 
|  |    686 | patterns.  A pattern can be any term (even
 | 
|  |    687 | a single identifier).  It may contain ``\texttt{\_}'', a wildcard standing
 | 
|  |    688 | for any term. Here are some
 | 
| 16519 |    689 | examples:
 | 
|  |    690 | \begin{ttbox}
 | 
|  |    691 | length
 | 
|  |    692 | "_ # _ = _ # _"
 | 
|  |    693 | "_ + _"
 | 
|  |    694 | "_ * (_ - (_::nat))"
 | 
|  |    695 | \end{ttbox}
 | 
| 16560 |    696 | Specifying types, as shown in the last example, 
 | 
|  |    697 | constrains searches involving overloaded operators.
 | 
| 16519 |    698 | 
 | 
|  |    699 | \begin{warn}
 | 
|  |    700 | Always use ``\texttt{\_}'' rather than variable names: searching for
 | 
|  |    701 | \texttt{"x + y"} will usually not find any matching theorems
 | 
| 16560 |    702 | because they would need to contain \texttt{x} and~\texttt{y} literally.
 | 
|  |    703 | When searching for infix operators, do not just type in the symbol,
 | 
|  |    704 | such as~\texttt{+}, but a proper term such as \texttt{"_ + _"}.
 | 
|  |    705 | This remark applies to more complicated syntaxes, too.
 | 
| 16519 |    706 | \end{warn}
 | 
|  |    707 | 
 | 
| 16560 |    708 | If you are looking for rewrite rules (possibly conditional) that could
 | 
|  |    709 | simplify some term, prefix the pattern with \texttt{simp:}.
 | 
| 16519 |    710 | \begin{ttbox}
 | 
|  |    711 | simp: "_ * (_ + _)"
 | 
|  |    712 | \end{ttbox}
 | 
| 16560 |    713 | This finds \emph{all} equations---not just those with a \isa{simp} attribute---whose conclusion has the form
 | 
| 16519 |    714 | \begin{isabelle}%
 | 
|  |    715 | \ \ \ \ \ {\isacharunderscore}\ {\isacharasterisk}\ {\isacharparenleft}{\isacharunderscore}\ {\isacharplus}\ {\isacharunderscore}{\isacharparenright}\ {\isacharequal}\ {\isasymdots}%
 | 
|  |    716 | \end{isabelle}
 | 
| 16560 |    717 | It only finds equations that can simplify the given pattern
 | 
|  |    718 | at the root, not somewhere inside: for example, equations of the form
 | 
|  |    719 | \isa{{\isacharunderscore}\ {\isacharplus}\ {\isacharunderscore}\ {\isacharequal}\ {\isasymdots}} do not match.
 | 
| 16519 |    720 | 
 | 
| 16560 |    721 | You may also search for theorems by name---you merely
 | 
|  |    722 | need to specify a substring. For example, you could search for all
 | 
| 16519 |    723 | commutativity theorems like this:
 | 
|  |    724 | \begin{ttbox}
 | 
|  |    725 | name: comm
 | 
|  |    726 | \end{ttbox}
 | 
|  |    727 | This retrieves all theorems whose name contains \texttt{comm}.
 | 
|  |    728 | 
 | 
| 16560 |    729 | Search criteria can also be negated by prefixing them with ``\texttt{-}''.
 | 
|  |    730 | For example,
 | 
| 16519 |    731 | \begin{ttbox}
 | 
| 16523 |    732 | -name: List
 | 
| 16519 |    733 | \end{ttbox}
 | 
| 16560 |    734 | finds theorems whose name does not contain \texttt{List}. You can use this
 | 
|  |    735 | to exclude particular theories from the search: the long name of
 | 
| 16519 |    736 | a theorem contains the name of the theory it comes from.
 | 
|  |    737 | 
 | 
| 16560 |    738 | Finallly, different search criteria can be combined arbitrarily. 
 | 
|  |    739 | The effect is conjuctive: Find returns the theorerms that satisfy all of
 | 
|  |    740 | the criteria. For example,
 | 
| 16519 |    741 | \begin{ttbox}
 | 
|  |    742 | "_ + _"  -"_ - _"  -simp: "_ * (_ + _)"  name: assoc
 | 
|  |    743 | \end{ttbox}
 | 
| 16560 |    744 | looks for theorems containing plus but not minus, and which do not simplify
 | 
|  |    745 | \mbox{\isa{{\isacharunderscore}\ {\isacharasterisk}\ {\isacharparenleft}{\isacharunderscore}\ {\isacharplus}\ {\isacharunderscore}{\isacharparenright}}} at the root, and whose name contains \texttt{assoc}.
 | 
| 16523 |    746 | 
 | 
| 16545 |    747 | Further search criteria are explained in \S\ref{sec:find2}.
 | 
| 16519 |    748 | 
 | 
| 16523 |    749 | \begin{pgnote}
 | 
|  |    750 | Proof General keeps a history of all your search expressions.
 | 
|  |    751 | If you click on \pgmenu{Find}, you can use the arrow keys to scroll
 | 
|  |    752 | through previous searches and just modify them. This saves you having
 | 
|  |    753 | to type in lengthy expressions again and again.
 | 
|  |    754 | \end{pgnote}%
 | 
| 16519 |    755 | \end{isamarkuptext}%
 | 
| 17056 |    756 | %
 | 
|  |    757 | \isadelimtheory
 | 
|  |    758 | %
 | 
|  |    759 | \endisadelimtheory
 | 
|  |    760 | %
 | 
|  |    761 | \isatagtheory
 | 
|  |    762 | %
 | 
|  |    763 | \endisatagtheory
 | 
|  |    764 | {\isafoldtheory}%
 | 
|  |    765 | %
 | 
|  |    766 | \isadelimtheory
 | 
|  |    767 | %
 | 
|  |    768 | \endisadelimtheory
 | 
| 9924 |    769 | \end{isabellebody}%
 | 
|  |    770 | %%% Local Variables:
 | 
|  |    771 | %%% mode: latex
 | 
|  |    772 | %%% TeX-master: "root"
 | 
|  |    773 | %%% End:
 |