doc-src/TutorialI/Misc/simp.thy
author nipkow
Mon, 19 Mar 2001 13:28:06 +0100
changeset 11215 b44ad7e4c4d2
parent 11214 3b3cc0cf533f
child 11309 d666f11ca2d4
permissions -rw-r--r--
*** empty log message ***
nipkow@9932
     1
(*<*)
nipkow@9932
     2
theory simp = Main:
nipkow@9932
     3
(*>*)
wenzelm@9922
     4
nipkow@11214
     5
subsection{*Simplification Rules*}
wenzelm@9922
     6
nipkow@9932
     7
text{*\indexbold{simplification rule}
nipkow@9932
     8
To facilitate simplification, theorems can be declared to be simplification
paulson@10795
     9
rules (by the attribute @{text"[simp]"}\index{*simp
nipkow@9932
    10
  (attribute)}), in which case proofs by simplification make use of these
nipkow@9932
    11
rules automatically. In addition the constructs \isacommand{datatype} and
nipkow@9932
    12
\isacommand{primrec} (and a few others) invisibly declare useful
nipkow@9932
    13
simplification rules. Explicit definitions are \emph{not} declared
nipkow@9932
    14
simplification rules automatically!
nipkow@9932
    15
nipkow@9932
    16
Not merely equations but pretty much any theorem can become a simplification
nipkow@9932
    17
rule. The simplifier will try to make sense of it.  For example, a theorem
nipkow@9932
    18
@{prop"~P"} is automatically turned into @{prop"P = False"}. The details
nipkow@9932
    19
are explained in \S\ref{sec:SimpHow}.
nipkow@9932
    20
nipkow@9932
    21
The simplification attribute of theorems can be turned on and off as follows:
nipkow@9932
    22
\begin{quote}
nipkow@9932
    23
\isacommand{declare} \textit{theorem-name}@{text"[simp]"}\\
nipkow@9932
    24
\isacommand{declare} \textit{theorem-name}@{text"[simp del]"}
nipkow@9932
    25
\end{quote}
nipkow@9932
    26
As a rule of thumb, equations that really simplify (like @{prop"rev(rev xs) =
nipkow@9932
    27
 xs"} and @{prop"xs @ [] = xs"}) should be made simplification
nipkow@9932
    28
rules.  Those of a more specific nature (e.g.\ distributivity laws, which
nipkow@9932
    29
alter the structure of terms considerably) should only be used selectively,
nipkow@9932
    30
i.e.\ they should not be default simplification rules.  Conversely, it may
nipkow@9932
    31
also happen that a simplification rule needs to be disabled in certain
nipkow@9932
    32
proofs.  Frequent changes in the simplification status of a theorem may
nipkow@9932
    33
indicate a badly designed theory.
nipkow@9932
    34
\begin{warn}
paulson@10795
    35
  Simplification may run forever, for example if both $f(x) = g(x)$ and
nipkow@9932
    36
  $g(x) = f(x)$ are simplification rules. It is the user's responsibility not
nipkow@9932
    37
  to include simplification rules that can lead to nontermination, either on
nipkow@9932
    38
  their own or in combination with other simplification rules.
nipkow@9932
    39
\end{warn}
nipkow@9932
    40
*}
nipkow@9932
    41
nipkow@11214
    42
subsection{*The Simplification Method*}
nipkow@9932
    43
nipkow@9932
    44
text{*\index{*simp (method)|bold}
nipkow@9932
    45
The general format of the simplification method is
nipkow@9932
    46
\begin{quote}
nipkow@9932
    47
@{text simp} \textit{list of modifiers}
nipkow@9932
    48
\end{quote}
paulson@10795
    49
where the list of \emph{modifiers} fine tunes the behaviour and may
nipkow@9932
    50
be empty. Most if not all of the proofs seen so far could have been performed
nipkow@9932
    51
with @{text simp} instead of \isa{auto}, except that @{text simp} attacks
nipkow@10971
    52
only the first subgoal and may thus need to be repeated --- use
nipkow@9932
    53
\isaindex{simp_all} to simplify all subgoals.
nipkow@9932
    54
Note that @{text simp} fails if nothing changes.
nipkow@9932
    55
*}
nipkow@9932
    56
nipkow@11214
    57
subsection{*Adding and Deleting Simplification Rules*}
nipkow@9932
    58
nipkow@9932
    59
text{*
nipkow@9932
    60
If a certain theorem is merely needed in a few proofs by simplification,
nipkow@9932
    61
we do not need to make it a global simplification rule. Instead we can modify
nipkow@9932
    62
the set of simplification rules used in a simplification step by adding rules
nipkow@9932
    63
to it and/or deleting rules from it. The two modifiers for this are
nipkow@9932
    64
\begin{quote}
nipkow@9932
    65
@{text"add:"} \textit{list of theorem names}\\
nipkow@9932
    66
@{text"del:"} \textit{list of theorem names}
nipkow@9932
    67
\end{quote}
nipkow@9932
    68
In case you want to use only a specific list of theorems and ignore all
nipkow@9932
    69
others:
nipkow@9932
    70
\begin{quote}
nipkow@9932
    71
@{text"only:"} \textit{list of theorem names}
nipkow@9932
    72
\end{quote}
nipkow@9932
    73
*}
nipkow@9932
    74
nipkow@11214
    75
subsection{*Assumptions*}
nipkow@9932
    76
nipkow@9932
    77
text{*\index{simplification!with/of assumptions}
nipkow@9932
    78
By default, assumptions are part of the simplification process: they are used
nipkow@9932
    79
as simplification rules and are simplified themselves. For example:
nipkow@9932
    80
*}
nipkow@9932
    81
nipkow@10171
    82
lemma "\<lbrakk> xs @ zs = ys @ xs; [] @ xs = [] @ [] \<rbrakk> \<Longrightarrow> ys = zs";
nipkow@10171
    83
apply simp;
nipkow@10171
    84
done
nipkow@9932
    85
nipkow@9932
    86
text{*\noindent
nipkow@9932
    87
The second assumption simplifies to @{term"xs = []"}, which in turn
nipkow@9932
    88
simplifies the first assumption to @{term"zs = ys"}, thus reducing the
nipkow@9932
    89
conclusion to @{term"ys = ys"} and hence to @{term"True"}.
nipkow@9932
    90
nipkow@9932
    91
In some cases this may be too much of a good thing and may lead to
nipkow@9932
    92
nontermination:
nipkow@9932
    93
*}
nipkow@9932
    94
nipkow@10171
    95
lemma "\<forall>x. f x = g (f (g x)) \<Longrightarrow> f [] = f [] @ []";
nipkow@9932
    96
nipkow@9932
    97
txt{*\noindent
nipkow@9932
    98
cannot be solved by an unmodified application of @{text"simp"} because the
nipkow@9932
    99
simplification rule @{term"f x = g (f (g x))"} extracted from the assumption
nipkow@9932
   100
does not terminate. Isabelle notices certain simple forms of
nipkow@9932
   101
nontermination but not this one. The problem can be circumvented by
nipkow@9932
   102
explicitly telling the simplifier to ignore the assumptions:
nipkow@9932
   103
*}
nipkow@9932
   104
nipkow@10171
   105
apply(simp (no_asm));
nipkow@10171
   106
done
nipkow@9932
   107
nipkow@9932
   108
text{*\noindent
nipkow@10971
   109
There are three modifiers that influence the treatment of assumptions:
nipkow@9932
   110
\begin{description}
nipkow@9932
   111
\item[@{text"(no_asm)"}]\indexbold{*no_asm}
nipkow@9932
   112
 means that assumptions are completely ignored.
nipkow@9932
   113
\item[@{text"(no_asm_simp)"}]\indexbold{*no_asm_simp}
nipkow@9932
   114
 means that the assumptions are not simplified but
nipkow@9932
   115
  are used in the simplification of the conclusion.
nipkow@9932
   116
\item[@{text"(no_asm_use)"}]\indexbold{*no_asm_use}
nipkow@9932
   117
 means that the assumptions are simplified but are not
nipkow@9932
   118
  used in the simplification of each other or the conclusion.
nipkow@9932
   119
\end{description}
paulson@10795
   120
Both @{text"(no_asm_simp)"} and @{text"(no_asm_use)"} run forever on
paulson@10795
   121
the problematic subgoal above.
nipkow@10971
   122
Note that only one of the modifiers is allowed, and it must precede all
nipkow@9932
   123
other arguments.
nipkow@11206
   124
nipkow@11206
   125
\begin{warn}
nipkow@11206
   126
Assumptions are simplified in a left-to-right fashion. If an
nipkow@11206
   127
assumption can help in simplifying one to the left of it, this may get
nipkow@11206
   128
overlooked. In such cases you have to rotate the assumptions explicitly:
nipkow@11206
   129
\isacommand{apply}@{text"(rotate_tac"}~$n$@{text")"}\indexbold{*rotate_tac}
nipkow@11206
   130
causes a cyclic shift by $n$ positions from right to left, if $n$ is
nipkow@11206
   131
positive, and from left to right, if $n$ is negative.
nipkow@11206
   132
Beware that such rotations make proofs quite brittle.
nipkow@11206
   133
\end{warn}
nipkow@9932
   134
*}
nipkow@9932
   135
nipkow@11214
   136
subsection{*Rewriting with Definitions*}
nipkow@9932
   137
nipkow@11215
   138
text{*\label{sec:Simp-with-Defs}\index{simplification!with definitions}
nipkow@9932
   139
Constant definitions (\S\ref{sec:ConstDefinitions}) can
nipkow@9932
   140
be used as simplification rules, but by default they are not.  Hence the
nipkow@9932
   141
simplifier does not expand them automatically, just as it should be:
nipkow@9932
   142
definitions are introduced for the purpose of abbreviating complex
nipkow@9932
   143
concepts. Of course we need to expand the definitions initially to derive
nipkow@9932
   144
enough lemmas that characterize the concept sufficiently for us to forget the
nipkow@9932
   145
original definition. For example, given
nipkow@9932
   146
*}
nipkow@9932
   147
nipkow@10788
   148
constdefs xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"
nipkow@10788
   149
         "xor A B \<equiv> (A \<and> \<not>B) \<or> (\<not>A \<and> B)";
nipkow@9932
   150
nipkow@9932
   151
text{*\noindent
nipkow@9932
   152
we may want to prove
nipkow@9932
   153
*}
nipkow@9932
   154
nipkow@10788
   155
lemma "xor A (\<not>A)";
nipkow@9932
   156
nipkow@9932
   157
txt{*\noindent
nipkow@9932
   158
Typically, the opening move consists in \emph{unfolding} the definition(s), which we need to
nipkow@9932
   159
get started, but nothing else:\indexbold{*unfold}\indexbold{definition!unfolding}
nipkow@9932
   160
*}
nipkow@9932
   161
nipkow@10788
   162
apply(simp only:xor_def);
nipkow@9932
   163
nipkow@9932
   164
txt{*\noindent
nipkow@9932
   165
In this particular case, the resulting goal
nipkow@10362
   166
@{subgoals[display,indent=0]}
nipkow@10171
   167
can be proved by simplification. Thus we could have proved the lemma outright by
nipkow@10788
   168
*}(*<*)oops;lemma "xor A (\<not>A)";(*>*)
nipkow@10788
   169
apply(simp add: xor_def)
nipkow@10171
   170
(*<*)done(*>*)
nipkow@9932
   171
text{*\noindent
nipkow@9932
   172
Of course we can also unfold definitions in the middle of a proof.
nipkow@9932
   173
nipkow@9932
   174
You should normally not turn a definition permanently into a simplification
nipkow@10983
   175
rule because this defeats the whole purpose.
nipkow@9932
   176
nipkow@9932
   177
\begin{warn}
nipkow@10971
   178
  If you have defined $f\,x\,y~\isasymequiv~t$ then you can only unfold
nipkow@10971
   179
  occurrences of $f$ with at least two arguments. This may be helpful for unfolding
nipkow@10971
   180
  $f$ selectively, but it may also get in the way. Defining
nipkow@10971
   181
  $f$~\isasymequiv~\isasymlambda$x\,y.\;t$ allows to unfold all occurrences of $f$.
nipkow@9932
   182
\end{warn}
nipkow@9932
   183
*}
nipkow@9932
   184
nipkow@11214
   185
subsection{*Simplifying {\tt\slshape let}-Expressions*}
nipkow@9932
   186
nipkow@9932
   187
text{*\index{simplification!of let-expressions}
nipkow@9932
   188
Proving a goal containing \isaindex{let}-expressions almost invariably
nipkow@9932
   189
requires the @{text"let"}-con\-structs to be expanded at some point. Since
paulson@10795
   190
@{text"let"}\ldots\isa{=}\ldots@{text"in"}{\ldots} is just syntactic sugar for a predefined constant
nipkow@9932
   191
(called @{term"Let"}), expanding @{text"let"}-constructs means rewriting with
nipkow@9932
   192
@{thm[source]Let_def}:
nipkow@9932
   193
*}
nipkow@9932
   194
nipkow@9932
   195
lemma "(let xs = [] in xs@ys@xs) = ys";
nipkow@10171
   196
apply(simp add: Let_def);
nipkow@10171
   197
done
nipkow@9932
   198
nipkow@9932
   199
text{*
nipkow@9932
   200
If, in a particular context, there is no danger of a combinatorial explosion
nipkow@9932
   201
of nested @{text"let"}s one could even simlify with @{thm[source]Let_def} by
nipkow@9932
   202
default:
nipkow@9932
   203
*}
nipkow@9932
   204
declare Let_def [simp]
nipkow@9932
   205
nipkow@11214
   206
subsection{*Conditional Equations*}
nipkow@9932
   207
nipkow@9932
   208
text{*
nipkow@9932
   209
So far all examples of rewrite rules were equations. The simplifier also
nipkow@9932
   210
accepts \emph{conditional} equations, for example
nipkow@9932
   211
*}
nipkow@9932
   212
nipkow@10171
   213
lemma hd_Cons_tl[simp]: "xs \<noteq> []  \<Longrightarrow>  hd xs # tl xs = xs";
nipkow@10171
   214
apply(case_tac xs, simp, simp);
nipkow@10171
   215
done
nipkow@9932
   216
nipkow@9932
   217
text{*\noindent
nipkow@9932
   218
Note the use of ``\ttindexboldpos{,}{$Isar}'' to string together a
nipkow@9932
   219
sequence of methods. Assuming that the simplification rule
nipkow@9932
   220
@{term"(rev xs = []) = (xs = [])"}
nipkow@9932
   221
is present as well,
paulson@10795
   222
the lemma below is proved by plain simplification:
nipkow@9932
   223
*}
nipkow@9932
   224
nipkow@10171
   225
lemma "xs \<noteq> [] \<Longrightarrow> hd(rev xs) # tl(rev xs) = rev xs";
nipkow@9932
   226
(*<*)
nipkow@9932
   227
by(simp);
nipkow@9932
   228
(*>*)
nipkow@9932
   229
text{*\noindent
paulson@10795
   230
The conditional equation @{thm[source]hd_Cons_tl} above
nipkow@9932
   231
can simplify @{term"hd(rev xs) # tl(rev xs)"} to @{term"rev xs"}
nipkow@9932
   232
because the corresponding precondition @{term"rev xs ~= []"}
nipkow@9932
   233
simplifies to @{term"xs ~= []"}, which is exactly the local
nipkow@9932
   234
assumption of the subgoal.
nipkow@9932
   235
*}
nipkow@9932
   236
nipkow@9932
   237
nipkow@11214
   238
subsection{*Automatic Case Splits*}
nipkow@9932
   239
nipkow@11215
   240
text{*\label{sec:AutoCaseSplits}\indexbold{case splits}\index{*split (method, attr.)|(}
nipkow@9932
   241
Goals containing @{text"if"}-expressions are usually proved by case
nipkow@9932
   242
distinction on the condition of the @{text"if"}. For example the goal
nipkow@9932
   243
*}
nipkow@9932
   244
nipkow@10171
   245
lemma "\<forall>xs. if xs = [] then rev xs = [] else rev xs \<noteq> []";
nipkow@9932
   246
nipkow@9932
   247
txt{*\noindent
nipkow@10654
   248
can be split by a special method @{text split}:
nipkow@9932
   249
*}
nipkow@9932
   250
nipkow@10654
   251
apply(split split_if)
nipkow@9932
   252
nipkow@10362
   253
txt{*\noindent
nipkow@10362
   254
@{subgoals[display,indent=0]}
nipkow@10654
   255
where \isaindexbold{split_if} is a theorem that expresses splitting of
nipkow@10654
   256
@{text"if"}s. Because
nipkow@9932
   257
case-splitting on @{text"if"}s is almost always the right proof strategy, the
nipkow@9932
   258
simplifier performs it automatically. Try \isacommand{apply}@{text"(simp)"}
nipkow@9932
   259
on the initial goal above.
nipkow@9932
   260
nipkow@9932
   261
This splitting idea generalizes from @{text"if"} to \isaindex{case}:
nipkow@10654
   262
*}(*<*)by simp(*>*)
nipkow@10171
   263
lemma "(case xs of [] \<Rightarrow> zs | y#ys \<Rightarrow> y#(ys@zs)) = xs@zs";
nipkow@10654
   264
apply(split list.split);
nipkow@9932
   265
nipkow@10362
   266
txt{*
nipkow@10362
   267
@{subgoals[display,indent=0]}
nipkow@9932
   268
In contrast to @{text"if"}-expressions, the simplifier does not split
nipkow@9932
   269
@{text"case"}-expressions by default because this can lead to nontermination
nipkow@10654
   270
in case of recursive datatypes. Therefore the simplifier has a modifier
nipkow@10654
   271
@{text split} for adding further splitting rules explicitly. This means the
nipkow@10654
   272
above lemma can be proved in one step by
nipkow@9932
   273
*}
nipkow@10362
   274
(*<*)oops;
nipkow@10171
   275
lemma "(case xs of [] \<Rightarrow> zs | y#ys \<Rightarrow> y#(ys@zs)) = xs@zs";
nipkow@9932
   276
(*>*)
nipkow@10171
   277
apply(simp split: list.split);
nipkow@10171
   278
(*<*)done(*>*)
nipkow@10654
   279
text{*\noindent
nipkow@10654
   280
whereas \isacommand{apply}@{text"(simp)"} alone will not succeed.
nipkow@9932
   281
nipkow@9932
   282
In general, every datatype $t$ comes with a theorem
nipkow@9932
   283
$t$@{text".split"} which can be declared to be a \bfindex{split rule} either
nipkow@9932
   284
locally as above, or by giving it the @{text"split"} attribute globally:
nipkow@9932
   285
*}
nipkow@9932
   286
nipkow@9932
   287
declare list.split [split]
nipkow@9932
   288
nipkow@9932
   289
text{*\noindent
nipkow@9932
   290
The @{text"split"} attribute can be removed with the @{text"del"} modifier,
nipkow@9932
   291
either locally
nipkow@9932
   292
*}
nipkow@9932
   293
(*<*)
nipkow@9932
   294
lemma "dummy=dummy";
nipkow@9932
   295
(*>*)
nipkow@9932
   296
apply(simp split del: split_if);
nipkow@9932
   297
(*<*)
nipkow@9932
   298
oops;
nipkow@9932
   299
(*>*)
nipkow@9932
   300
text{*\noindent
nipkow@9932
   301
or globally:
nipkow@9932
   302
*}
nipkow@9932
   303
declare list.split [split del]
nipkow@9932
   304
nipkow@9932
   305
text{*
nipkow@10654
   306
In polished proofs the @{text split} method is rarely used on its own
nipkow@10654
   307
but always as part of the simplifier. However, if a goal contains
nipkow@10654
   308
multiple splittable constructs, the @{text split} method can be
nipkow@10654
   309
helpful in selectively exploring the effects of splitting.
nipkow@10654
   310
nipkow@9932
   311
The above split rules intentionally only affect the conclusion of a
nipkow@9932
   312
subgoal.  If you want to split an @{text"if"} or @{text"case"}-expression in
nipkow@9932
   313
the assumptions, you have to apply @{thm[source]split_if_asm} or
nipkow@9932
   314
$t$@{text".split_asm"}:
nipkow@9932
   315
*}
nipkow@9932
   316
nipkow@10654
   317
lemma "if xs = [] then ys \<noteq> [] else ys = [] \<Longrightarrow> xs @ ys \<noteq> []"
nipkow@10654
   318
apply(split split_if_asm)
nipkow@9932
   319
nipkow@10362
   320
txt{*\noindent
nipkow@9932
   321
In contrast to splitting the conclusion, this actually creates two
nipkow@9932
   322
separate subgoals (which are solved by @{text"simp_all"}):
nipkow@10362
   323
@{subgoals[display,indent=0]}
nipkow@9932
   324
If you need to split both in the assumptions and the conclusion,
nipkow@9932
   325
use $t$@{text".splits"} which subsumes $t$@{text".split"} and
nipkow@9932
   326
$t$@{text".split_asm"}. Analogously, there is @{thm[source]if_splits}.
nipkow@9932
   327
nipkow@9932
   328
\begin{warn}
nipkow@9932
   329
  The simplifier merely simplifies the condition of an \isa{if} but not the
nipkow@9932
   330
  \isa{then} or \isa{else} parts. The latter are simplified only after the
nipkow@9932
   331
  condition reduces to \isa{True} or \isa{False}, or after splitting. The
nipkow@9932
   332
  same is true for \isaindex{case}-expressions: only the selector is
nipkow@9932
   333
  simplified at first, until either the expression reduces to one of the
nipkow@9932
   334
  cases or it is split.
nipkow@10654
   335
\end{warn}\index{*split (method, attr.)|)}
nipkow@9932
   336
*}
nipkow@10362
   337
(*<*)
nipkow@10362
   338
by(simp_all)
nipkow@10362
   339
(*>*)
nipkow@9932
   340
nipkow@11214
   341
subsection{*Arithmetic*}
nipkow@9932
   342
nipkow@9932
   343
text{*\index{arithmetic}
nipkow@9932
   344
The simplifier routinely solves a small class of linear arithmetic formulae
nipkow@9932
   345
(over type \isa{nat} and other numeric types): it only takes into account
paulson@10795
   346
assumptions and conclusions that are relations
paulson@10795
   347
($=$, $\le$, $<$, possibly negated) and it only knows about addition. Thus
nipkow@9932
   348
*}
nipkow@9932
   349
nipkow@10171
   350
lemma "\<lbrakk> \<not> m < n; m < n+1 \<rbrakk> \<Longrightarrow> m = n"
nipkow@9932
   351
(*<*)by(auto)(*>*)
nipkow@9932
   352
nipkow@9932
   353
text{*\noindent
nipkow@9932
   354
is proved by simplification, whereas the only slightly more complex
nipkow@9932
   355
*}
nipkow@9932
   356
nipkow@10171
   357
lemma "\<not> m < n \<and> m < n+1 \<Longrightarrow> m = n";
nipkow@9932
   358
(*<*)by(arith)(*>*)
nipkow@9932
   359
nipkow@9932
   360
text{*\noindent
nipkow@9932
   361
is not proved by simplification and requires @{text arith}.
nipkow@9932
   362
*}
nipkow@9932
   363
nipkow@9932
   364
nipkow@11214
   365
subsection{*Tracing*}
nipkow@9932
   366
text{*\indexbold{tracing the simplifier}
nipkow@9932
   367
Using the simplifier effectively may take a bit of experimentation.  Set the
nipkow@9932
   368
\isaindexbold{trace_simp} \rmindex{flag} to get a better idea of what is going
nipkow@9932
   369
on:
nipkow@9932
   370
*}
nipkow@9932
   371
nipkow@9932
   372
ML "set trace_simp";
nipkow@9932
   373
lemma "rev [a] = []";
nipkow@9932
   374
apply(simp);
nipkow@9932
   375
(*<*)oops(*>*)
nipkow@9932
   376
nipkow@9932
   377
text{*\noindent
nipkow@9932
   378
produces the trace
nipkow@9932
   379
nipkow@9932
   380
\begin{ttbox}\makeatother
nipkow@9932
   381
Applying instance of rewrite rule:
nipkow@9932
   382
rev (?x1 \# ?xs1) == rev ?xs1 @ [?x1]
nipkow@9932
   383
Rewriting:
nipkow@10971
   384
rev [a] == rev [] @ [a]
nipkow@9932
   385
Applying instance of rewrite rule:
nipkow@9932
   386
rev [] == []
nipkow@9932
   387
Rewriting:
nipkow@9932
   388
rev [] == []
nipkow@9932
   389
Applying instance of rewrite rule:
nipkow@9932
   390
[] @ ?y == ?y
nipkow@9932
   391
Rewriting:
nipkow@10971
   392
[] @ [a] == [a]
nipkow@9932
   393
Applying instance of rewrite rule:
nipkow@9932
   394
?x3 \# ?t3 = ?t3 == False
nipkow@9932
   395
Rewriting:
nipkow@10971
   396
[a] = [] == False
nipkow@9932
   397
\end{ttbox}
nipkow@9932
   398
nipkow@9932
   399
In more complicated cases, the trace can be quite lenghty, especially since
nipkow@9932
   400
invocations of the simplifier are often nested (e.g.\ when solving conditions
nipkow@9932
   401
of rewrite rules). Thus it is advisable to reset it:
nipkow@9932
   402
*}
nipkow@9932
   403
nipkow@9932
   404
ML "reset trace_simp";
nipkow@9932
   405
nipkow@9932
   406
(*<*)
wenzelm@9922
   407
end
nipkow@9932
   408
(*>*)