doc-src/TutorialI/Misc/simp.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 12 Aug 2010 15:03:34 +0200
branchisac-from-Isabelle2009-2
changeset 37913 20e3616b2d9c
parent 36069 a15454b23ebd
child 41136 7695e4de4d86
permissions -rw-r--r--
prepare reactivation of isac-update-Isa09-2
nipkow@9932
     1
(*<*)
haftmann@16417
     2
theory simp imports Main begin
nipkow@9932
     3
(*>*)
wenzelm@9922
     4
nipkow@11214
     5
subsection{*Simplification Rules*}
wenzelm@9922
     6
paulson@11458
     7
text{*\index{simplification rules}
paulson@11458
     8
To facilitate simplification,  
paulson@11458
     9
the attribute @{text"[simp]"}\index{*simp (attribute)}
paulson@11458
    10
declares theorems to be simplification rules, which the simplifier
paulson@11458
    11
will use automatically.  In addition, \isacommand{datatype} and
paulson@11458
    12
\isacommand{primrec} declarations (and a few others) 
paulson@11458
    13
implicitly declare some simplification rules.  
paulson@11458
    14
Explicit definitions are \emph{not} declared as 
nipkow@9932
    15
simplification rules automatically!
nipkow@9932
    16
paulson@11458
    17
Nearly any theorem can become a simplification
paulson@11458
    18
rule. The simplifier will try to transform it into an equation.  
paulson@11458
    19
For example, the theorem
paulson@11458
    20
@{prop"~P"} is turned into @{prop"P = False"}. The details
nipkow@9932
    21
are explained in \S\ref{sec:SimpHow}.
nipkow@9932
    22
paulson@11458
    23
The simplification attribute of theorems can be turned on and off:%
nipkow@12489
    24
\index{*simp del (attribute)}
nipkow@9932
    25
\begin{quote}
nipkow@9932
    26
\isacommand{declare} \textit{theorem-name}@{text"[simp]"}\\
nipkow@9932
    27
\isacommand{declare} \textit{theorem-name}@{text"[simp del]"}
nipkow@9932
    28
\end{quote}
paulson@11309
    29
Only equations that really simplify, like \isa{rev\
paulson@11309
    30
{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs} and
paulson@11309
    31
\isa{xs\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}\
paulson@11309
    32
{\isacharequal}\ xs}, should be declared as default simplification rules. 
paulson@11309
    33
More specific ones should only be used selectively and should
paulson@11309
    34
not be made default.  Distributivity laws, for example, alter
paulson@11309
    35
the structure of terms and can produce an exponential blow-up instead of
paulson@11309
    36
simplification.  A default simplification rule may
paulson@11309
    37
need to be disabled in certain proofs.  Frequent changes in the simplification
paulson@11309
    38
status of a theorem may indicate an unwise use of defaults.
nipkow@9932
    39
\begin{warn}
paulson@11458
    40
  Simplification can run forever, for example if both $f(x) = g(x)$ and
nipkow@9932
    41
  $g(x) = f(x)$ are simplification rules. It is the user's responsibility not
nipkow@9932
    42
  to include simplification rules that can lead to nontermination, either on
nipkow@9932
    43
  their own or in combination with other simplification rules.
nipkow@9932
    44
\end{warn}
nipkow@12332
    45
\begin{warn}
nipkow@12332
    46
  It is inadvisable to toggle the simplification attribute of a
nipkow@12332
    47
  theorem from a parent theory $A$ in a child theory $B$ for good.
nipkow@12332
    48
  The reason is that if some theory $C$ is based both on $B$ and (via a
nipkow@13814
    49
  different path) on $A$, it is not defined what the simplification attribute
nipkow@12332
    50
  of that theorem will be in $C$: it could be either.
nipkow@12332
    51
\end{warn}
paulson@11458
    52
*} 
nipkow@9932
    53
paulson@11458
    54
subsection{*The {\tt\slshape simp}  Method*}
nipkow@9932
    55
nipkow@9932
    56
text{*\index{*simp (method)|bold}
nipkow@9932
    57
The general format of the simplification method is
nipkow@9932
    58
\begin{quote}
nipkow@9932
    59
@{text simp} \textit{list of modifiers}
nipkow@9932
    60
\end{quote}
paulson@10795
    61
where the list of \emph{modifiers} fine tunes the behaviour and may
paulson@11309
    62
be empty. Specific modifiers are discussed below.  Most if not all of the
paulson@11309
    63
proofs seen so far could have been performed
nipkow@9932
    64
with @{text simp} instead of \isa{auto}, except that @{text simp} attacks
nipkow@10971
    65
only the first subgoal and may thus need to be repeated --- use
paulson@11428
    66
\methdx{simp_all} to simplify all subgoals.
paulson@11458
    67
If nothing changes, @{text simp} fails.
nipkow@9932
    68
*}
nipkow@9932
    69
nipkow@11214
    70
subsection{*Adding and Deleting Simplification Rules*}
nipkow@9932
    71
nipkow@9932
    72
text{*
paulson@11458
    73
\index{simplification rules!adding and deleting}%
nipkow@9932
    74
If a certain theorem is merely needed in a few proofs by simplification,
nipkow@9932
    75
we do not need to make it a global simplification rule. Instead we can modify
nipkow@9932
    76
the set of simplification rules used in a simplification step by adding rules
nipkow@9932
    77
to it and/or deleting rules from it. The two modifiers for this are
nipkow@9932
    78
\begin{quote}
paulson@11458
    79
@{text"add:"} \textit{list of theorem names}\index{*add (modifier)}\\
paulson@11458
    80
@{text"del:"} \textit{list of theorem names}\index{*del (modifier)}
nipkow@9932
    81
\end{quote}
paulson@11458
    82
Or you can use a specific list of theorems and omit all others:
nipkow@9932
    83
\begin{quote}
paulson@11458
    84
@{text"only:"} \textit{list of theorem names}\index{*only (modifier)}
nipkow@9932
    85
\end{quote}
paulson@11309
    86
In this example, we invoke the simplifier, adding two distributive
paulson@11309
    87
laws:
paulson@11309
    88
\begin{quote}
paulson@11309
    89
\isacommand{apply}@{text"(simp add: mod_mult_distrib add_mult_distrib)"}
paulson@11309
    90
\end{quote}
nipkow@9932
    91
*}
nipkow@9932
    92
nipkow@11214
    93
subsection{*Assumptions*}
nipkow@9932
    94
nipkow@9932
    95
text{*\index{simplification!with/of assumptions}
nipkow@9932
    96
By default, assumptions are part of the simplification process: they are used
nipkow@9932
    97
as simplification rules and are simplified themselves. For example:
nipkow@9932
    98
*}
nipkow@9932
    99
wenzelm@12631
   100
lemma "\<lbrakk> xs @ zs = ys @ xs; [] @ xs = [] @ [] \<rbrakk> \<Longrightarrow> ys = zs"
wenzelm@12631
   101
apply simp
nipkow@10171
   102
done
nipkow@9932
   103
nipkow@9932
   104
text{*\noindent
nipkow@9932
   105
The second assumption simplifies to @{term"xs = []"}, which in turn
nipkow@9932
   106
simplifies the first assumption to @{term"zs = ys"}, thus reducing the
nipkow@9932
   107
conclusion to @{term"ys = ys"} and hence to @{term"True"}.
nipkow@9932
   108
paulson@11458
   109
In some cases, using the assumptions can lead to nontermination:
nipkow@9932
   110
*}
nipkow@9932
   111
wenzelm@12631
   112
lemma "\<forall>x. f x = g (f (g x)) \<Longrightarrow> f [] = f [] @ []"
nipkow@9932
   113
nipkow@9932
   114
txt{*\noindent
paulson@11458
   115
An unmodified application of @{text"simp"} loops.  The culprit is the
paulson@11458
   116
simplification rule @{term"f x = g (f (g x))"}, which is extracted from
paulson@11458
   117
the assumption.  (Isabelle notices certain simple forms of
paulson@11458
   118
nontermination but not this one.)  The problem can be circumvented by
paulson@11458
   119
telling the simplifier to ignore the assumptions:
nipkow@9932
   120
*}
nipkow@9932
   121
wenzelm@12631
   122
apply(simp (no_asm))
nipkow@10171
   123
done
nipkow@9932
   124
nipkow@9932
   125
text{*\noindent
paulson@11458
   126
Three modifiers influence the treatment of assumptions:
nipkow@9932
   127
\begin{description}
paulson@11458
   128
\item[@{text"(no_asm)"}]\index{*no_asm (modifier)}
nipkow@9932
   129
 means that assumptions are completely ignored.
paulson@11458
   130
\item[@{text"(no_asm_simp)"}]\index{*no_asm_simp (modifier)}
nipkow@9932
   131
 means that the assumptions are not simplified but
nipkow@9932
   132
  are used in the simplification of the conclusion.
paulson@11458
   133
\item[@{text"(no_asm_use)"}]\index{*no_asm_use (modifier)}
nipkow@9932
   134
 means that the assumptions are simplified but are not
nipkow@9932
   135
  used in the simplification of each other or the conclusion.
nipkow@9932
   136
\end{description}
paulson@11458
   137
Only one of the modifiers is allowed, and it must precede all
paulson@11309
   138
other modifiers.
nipkow@13623
   139
%\begin{warn}
nipkow@13623
   140
%Assumptions are simplified in a left-to-right fashion. If an
nipkow@13623
   141
%assumption can help in simplifying one to the left of it, this may get
nipkow@13623
   142
%overlooked. In such cases you have to rotate the assumptions explicitly:
nipkow@13623
   143
%\isacommand{apply}@ {text"("}\methdx{rotate_tac}~$n$@ {text")"}
nipkow@13623
   144
%causes a cyclic shift by $n$ positions from right to left, if $n$ is
nipkow@13623
   145
%positive, and from left to right, if $n$ is negative.
nipkow@13623
   146
%Beware that such rotations make proofs quite brittle.
nipkow@13623
   147
%\end{warn}
nipkow@9932
   148
*}
nipkow@9932
   149
nipkow@11214
   150
subsection{*Rewriting with Definitions*}
nipkow@9932
   151
nipkow@11215
   152
text{*\label{sec:Simp-with-Defs}\index{simplification!with definitions}
paulson@11458
   153
Constant definitions (\S\ref{sec:ConstDefinitions}) can be used as
paulson@11458
   154
simplification rules, but by default they are not: the simplifier does not
paulson@11458
   155
expand them automatically.  Definitions are intended for introducing abstract
wenzelm@12582
   156
concepts and not merely as abbreviations.  Of course, we need to expand
paulson@11458
   157
the definition initially, but once we have proved enough abstract properties
paulson@11458
   158
of the new constant, we can forget its original definition.  This style makes
paulson@11458
   159
proofs more robust: if the definition has to be changed,
paulson@11458
   160
only the proofs of the abstract properties will be affected.
paulson@11458
   161
paulson@11458
   162
For example, given *}
nipkow@9932
   163
nipkow@27027
   164
definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
nipkow@27027
   165
"xor A B \<equiv> (A \<and> \<not>B) \<or> (\<not>A \<and> B)"
nipkow@9932
   166
nipkow@9932
   167
text{*\noindent
nipkow@9932
   168
we may want to prove
nipkow@9932
   169
*}
nipkow@9932
   170
wenzelm@12631
   171
lemma "xor A (\<not>A)"
nipkow@9932
   172
nipkow@9932
   173
txt{*\noindent
paulson@11428
   174
Typically, we begin by unfolding some definitions:
paulson@11428
   175
\indexbold{definitions!unfolding}
nipkow@9932
   176
*}
nipkow@9932
   177
wenzelm@12631
   178
apply(simp only: xor_def)
nipkow@9932
   179
nipkow@9932
   180
txt{*\noindent
nipkow@9932
   181
In this particular case, the resulting goal
nipkow@10362
   182
@{subgoals[display,indent=0]}
nipkow@10171
   183
can be proved by simplification. Thus we could have proved the lemma outright by
wenzelm@12631
   184
*}(*<*)oops lemma "xor A (\<not>A)"(*>*)
nipkow@10788
   185
apply(simp add: xor_def)
nipkow@10171
   186
(*<*)done(*>*)
nipkow@9932
   187
text{*\noindent
nipkow@9932
   188
Of course we can also unfold definitions in the middle of a proof.
nipkow@9932
   189
nipkow@9932
   190
\begin{warn}
nipkow@10971
   191
  If you have defined $f\,x\,y~\isasymequiv~t$ then you can only unfold
nipkow@10971
   192
  occurrences of $f$ with at least two arguments. This may be helpful for unfolding
nipkow@10971
   193
  $f$ selectively, but it may also get in the way. Defining
nipkow@10971
   194
  $f$~\isasymequiv~\isasymlambda$x\,y.\;t$ allows to unfold all occurrences of $f$.
nipkow@9932
   195
\end{warn}
nipkow@12473
   196
nipkow@12473
   197
There is also the special method \isa{unfold}\index{*unfold (method)|bold}
nipkow@12473
   198
which merely unfolds
nipkow@12473
   199
one or several definitions, as in \isacommand{apply}\isa{(unfold xor_def)}.
nipkow@12473
   200
This is can be useful in situations where \isa{simp} does too much.
nipkow@12533
   201
Warning: \isa{unfold} acts on all subgoals!
nipkow@9932
   202
*}
nipkow@9932
   203
nipkow@11214
   204
subsection{*Simplifying {\tt\slshape let}-Expressions*}
nipkow@9932
   205
paulson@11458
   206
text{*\index{simplification!of \isa{let}-expressions}\index{*let expressions}%
paulson@11458
   207
Proving a goal containing \isa{let}-expressions almost invariably requires the
paulson@11458
   208
@{text"let"}-con\-structs to be expanded at some point. Since
paulson@11458
   209
@{text"let"}\ldots\isa{=}\ldots@{text"in"}{\ldots} is just syntactic sugar for
paulson@11458
   210
the predefined constant @{term"Let"}, expanding @{text"let"}-constructs
paulson@11458
   211
means rewriting with \tdx{Let_def}: *}
nipkow@9932
   212
wenzelm@12631
   213
lemma "(let xs = [] in xs@ys@xs) = ys"
wenzelm@12631
   214
apply(simp add: Let_def)
nipkow@10171
   215
done
nipkow@9932
   216
nipkow@9932
   217
text{*
nipkow@9932
   218
If, in a particular context, there is no danger of a combinatorial explosion
paulson@11458
   219
of nested @{text"let"}s, you could even simplify with @{thm[source]Let_def} by
nipkow@9932
   220
default:
nipkow@9932
   221
*}
nipkow@9932
   222
declare Let_def [simp]
nipkow@9932
   223
paulson@11458
   224
subsection{*Conditional Simplification Rules*}
nipkow@9932
   225
nipkow@9932
   226
text{*
paulson@11458
   227
\index{conditional simplification rules}%
nipkow@9932
   228
So far all examples of rewrite rules were equations. The simplifier also
nipkow@9932
   229
accepts \emph{conditional} equations, for example
nipkow@9932
   230
*}
nipkow@9932
   231
wenzelm@12631
   232
lemma hd_Cons_tl[simp]: "xs \<noteq> []  \<Longrightarrow>  hd xs # tl xs = xs"
wenzelm@12631
   233
apply(case_tac xs, simp, simp)
nipkow@10171
   234
done
nipkow@9932
   235
nipkow@9932
   236
text{*\noindent
nipkow@9932
   237
Note the use of ``\ttindexboldpos{,}{$Isar}'' to string together a
nipkow@9932
   238
sequence of methods. Assuming that the simplification rule
nipkow@9932
   239
@{term"(rev xs = []) = (xs = [])"}
nipkow@9932
   240
is present as well,
paulson@10795
   241
the lemma below is proved by plain simplification:
nipkow@9932
   242
*}
nipkow@9932
   243
wenzelm@12631
   244
lemma "xs \<noteq> [] \<Longrightarrow> hd(rev xs) # tl(rev xs) = rev xs"
nipkow@9932
   245
(*<*)
wenzelm@12631
   246
by(simp)
nipkow@9932
   247
(*>*)
nipkow@9932
   248
text{*\noindent
paulson@10795
   249
The conditional equation @{thm[source]hd_Cons_tl} above
nipkow@9932
   250
can simplify @{term"hd(rev xs) # tl(rev xs)"} to @{term"rev xs"}
nipkow@9932
   251
because the corresponding precondition @{term"rev xs ~= []"}
nipkow@9932
   252
simplifies to @{term"xs ~= []"}, which is exactly the local
nipkow@9932
   253
assumption of the subgoal.
nipkow@9932
   254
*}
nipkow@9932
   255
nipkow@9932
   256
nipkow@11214
   257
subsection{*Automatic Case Splits*}
nipkow@9932
   258
paulson@11458
   259
text{*\label{sec:AutoCaseSplits}\indexbold{case splits}%
paulson@11458
   260
Goals containing @{text"if"}-expressions\index{*if expressions!splitting of}
paulson@11458
   261
are usually proved by case
paulson@11458
   262
distinction on the boolean condition.  Here is an example:
nipkow@9932
   263
*}
nipkow@9932
   264
wenzelm@12631
   265
lemma "\<forall>xs. if xs = [] then rev xs = [] else rev xs \<noteq> []"
nipkow@9932
   266
nipkow@9932
   267
txt{*\noindent
paulson@11458
   268
The goal can be split by a special method, \methdx{split}:
nipkow@9932
   269
*}
nipkow@9932
   270
nipkow@10654
   271
apply(split split_if)
nipkow@9932
   272
nipkow@10362
   273
txt{*\noindent
nipkow@10362
   274
@{subgoals[display,indent=0]}
paulson@11428
   275
where \tdx{split_if} is a theorem that expresses splitting of
nipkow@10654
   276
@{text"if"}s. Because
paulson@11458
   277
splitting the @{text"if"}s is usually the right proof strategy, the
paulson@11458
   278
simplifier does it automatically.  Try \isacommand{apply}@{text"(simp)"}
nipkow@9932
   279
on the initial goal above.
nipkow@9932
   280
paulson@11428
   281
This splitting idea generalizes from @{text"if"} to \sdx{case}.
paulson@11458
   282
Let us simplify a case analysis over lists:\index{*list.split (theorem)}
nipkow@10654
   283
*}(*<*)by simp(*>*)
wenzelm@12631
   284
lemma "(case xs of [] \<Rightarrow> zs | y#ys \<Rightarrow> y#(ys@zs)) = xs@zs"
wenzelm@12631
   285
apply(split list.split)
paulson@11309
   286
 
nipkow@10362
   287
txt{*
nipkow@10362
   288
@{subgoals[display,indent=0]}
paulson@11458
   289
The simplifier does not split
paulson@11458
   290
@{text"case"}-expressions, as it does @{text"if"}-expressions, 
paulson@11458
   291
because with recursive datatypes it could lead to nontermination.
paulson@11458
   292
Instead, the simplifier has a modifier
paulson@11494
   293
@{text split}\index{*split (modifier)} 
paulson@11458
   294
for adding splitting rules explicitly.  The
paulson@11458
   295
lemma above can be proved in one step by
nipkow@9932
   296
*}
wenzelm@12631
   297
(*<*)oops
wenzelm@12631
   298
lemma "(case xs of [] \<Rightarrow> zs | y#ys \<Rightarrow> y#(ys@zs)) = xs@zs"
nipkow@9932
   299
(*>*)
wenzelm@12631
   300
apply(simp split: list.split)
nipkow@10171
   301
(*<*)done(*>*)
nipkow@10654
   302
text{*\noindent
nipkow@10654
   303
whereas \isacommand{apply}@{text"(simp)"} alone will not succeed.
nipkow@9932
   304
paulson@11458
   305
Every datatype $t$ comes with a theorem
nipkow@9932
   306
$t$@{text".split"} which can be declared to be a \bfindex{split rule} either
paulson@11458
   307
locally as above, or by giving it the \attrdx{split} attribute globally:
nipkow@9932
   308
*}
nipkow@9932
   309
nipkow@9932
   310
declare list.split [split]
nipkow@9932
   311
nipkow@9932
   312
text{*\noindent
nipkow@9932
   313
The @{text"split"} attribute can be removed with the @{text"del"} modifier,
nipkow@9932
   314
either locally
nipkow@9932
   315
*}
nipkow@9932
   316
(*<*)
wenzelm@12631
   317
lemma "dummy=dummy"
nipkow@9932
   318
(*>*)
wenzelm@12631
   319
apply(simp split del: split_if)
nipkow@9932
   320
(*<*)
wenzelm@12631
   321
oops
nipkow@9932
   322
(*>*)
nipkow@9932
   323
text{*\noindent
nipkow@9932
   324
or globally:
nipkow@9932
   325
*}
nipkow@9932
   326
declare list.split [split del]
nipkow@9932
   327
nipkow@9932
   328
text{*
paulson@11458
   329
Polished proofs typically perform splitting within @{text simp} rather than 
paulson@11458
   330
invoking the @{text split} method.  However, if a goal contains
wenzelm@19792
   331
several @{text "if"} and @{text case} expressions, 
paulson@11458
   332
the @{text split} method can be
nipkow@10654
   333
helpful in selectively exploring the effects of splitting.
nipkow@10654
   334
paulson@11458
   335
The split rules shown above are intended to affect only the subgoal's
paulson@11458
   336
conclusion.  If you want to split an @{text"if"} or @{text"case"}-expression
paulson@11458
   337
in the assumptions, you have to apply \tdx{split_if_asm} or
paulson@11458
   338
$t$@{text".split_asm"}: *}
nipkow@9932
   339
nipkow@10654
   340
lemma "if xs = [] then ys \<noteq> [] else ys = [] \<Longrightarrow> xs @ ys \<noteq> []"
nipkow@10654
   341
apply(split split_if_asm)
nipkow@9932
   342
nipkow@10362
   343
txt{*\noindent
paulson@11458
   344
Unlike splitting the conclusion, this step creates two
paulson@11458
   345
separate subgoals, which here can be solved by @{text"simp_all"}:
nipkow@10362
   346
@{subgoals[display,indent=0]}
nipkow@9932
   347
If you need to split both in the assumptions and the conclusion,
nipkow@9932
   348
use $t$@{text".splits"} which subsumes $t$@{text".split"} and
nipkow@9932
   349
$t$@{text".split_asm"}. Analogously, there is @{thm[source]if_splits}.
nipkow@9932
   350
nipkow@9932
   351
\begin{warn}
paulson@11458
   352
  The simplifier merely simplifies the condition of an 
paulson@11458
   353
  \isa{if}\index{*if expressions!simplification of} but not the
nipkow@9932
   354
  \isa{then} or \isa{else} parts. The latter are simplified only after the
nipkow@9932
   355
  condition reduces to \isa{True} or \isa{False}, or after splitting. The
paulson@11428
   356
  same is true for \sdx{case}-expressions: only the selector is
nipkow@9932
   357
  simplified at first, until either the expression reduces to one of the
nipkow@9932
   358
  cases or it is split.
paulson@11458
   359
\end{warn}
nipkow@9932
   360
*}
nipkow@10362
   361
(*<*)
nipkow@10362
   362
by(simp_all)
nipkow@10362
   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@16523
   368
Proof General flag \pgmenu{Isabelle} $>$ \pgmenu{Settings} $>$ \pgmenu{Trace Simplifier} to get a better idea of what is going on:
nipkow@9932
   369
*}
nipkow@9932
   370
wenzelm@12631
   371
lemma "rev [a] = []"
wenzelm@12631
   372
apply(simp)
nipkow@9932
   373
(*<*)oops(*>*)
nipkow@9932
   374
nipkow@9932
   375
text{*\noindent
nipkow@16523
   376
produces the following trace in Proof General's \pgmenu{Trace} buffer:
nipkow@9932
   377
nipkow@9932
   378
\begin{ttbox}\makeatother
nipkow@16518
   379
[1]Applying instance of rewrite rule "List.rev.simps_2":
nipkow@16359
   380
rev (?x1 # ?xs1) \(\equiv\) rev ?xs1 @ [?x1]
nipkow@16359
   381
nipkow@16518
   382
[1]Rewriting:
nipkow@16359
   383
rev [a] \(\equiv\) rev [] @ [a]
nipkow@16359
   384
nipkow@16518
   385
[1]Applying instance of rewrite rule "List.rev.simps_1":
nipkow@16359
   386
rev [] \(\equiv\) []
nipkow@16359
   387
nipkow@16518
   388
[1]Rewriting:
nipkow@16359
   389
rev [] \(\equiv\) []
nipkow@16359
   390
nipkow@16518
   391
[1]Applying instance of rewrite rule "List.op @.append_Nil":
nipkow@16359
   392
[] @ ?y \(\equiv\) ?y
nipkow@16359
   393
nipkow@16518
   394
[1]Rewriting:
nipkow@16359
   395
[] @ [a] \(\equiv\) [a]
nipkow@16359
   396
nipkow@16518
   397
[1]Applying instance of rewrite rule
nipkow@16359
   398
?x2 # ?t1 = ?t1 \(\equiv\) False
nipkow@16359
   399
nipkow@16518
   400
[1]Rewriting:
nipkow@16359
   401
[a] = [] \(\equiv\) False
nipkow@9932
   402
\end{ttbox}
nipkow@16359
   403
The trace lists each rule being applied, both in its general form and
nipkow@16359
   404
the instance being used. The \texttt{[}$i$\texttt{]} in front (where
nipkow@16518
   405
above $i$ is always \texttt{1}) indicates that we are inside the $i$th
nipkow@16518
   406
invocation of the simplifier. Each attempt to apply a
nipkow@16359
   407
conditional rule shows the rule followed by the trace of the
nipkow@16359
   408
(recursive!) simplification of the conditions, the latter prefixed by
nipkow@16359
   409
\texttt{[}$i+1$\texttt{]} instead of \texttt{[}$i$\texttt{]}.
nipkow@16359
   410
Another source of recursive invocations of the simplifier are
nipkow@35992
   411
proofs of arithmetic formulae. By default, recursive invocations are not shown,
nipkow@35992
   412
you must increase the trace depth via \pgmenu{Isabelle} $>$ \pgmenu{Settings} $>$ \pgmenu{Trace Simplifier Depth}.
nipkow@9932
   413
nipkow@16359
   414
Many other hints about the simplifier's actions may appear.
paulson@11309
   415
nipkow@16359
   416
In more complicated cases, the trace can be very lengthy.  Thus it is
nipkow@16523
   417
advisable to reset the \pgmenu{Trace Simplifier} flag after having
nipkow@35992
   418
obtained the desired trace.
nipkow@36069
   419
Since this is easily forgotten (and may have the unpleasant effect of
nipkow@36069
   420
swamping the interface with trace information), here is how you can switch
nipkow@36069
   421
the trace on locally in a proof: *}
nipkow@36069
   422
nipkow@36069
   423
(*<*)lemma "x=x"
nipkow@36069
   424
(*>*)
nipkow@36069
   425
using [[trace_simp=true]]
nipkow@36069
   426
apply simp
nipkow@36069
   427
(*<*)oops(*>*)
nipkow@36069
   428
nipkow@36069
   429
text{* \noindent
nipkow@36069
   430
Within the current proof, all simplifications in subsequent proof steps
nipkow@36069
   431
will be traced, but the text reminds you to remove the \isa{using} clause
nipkow@36069
   432
after it has done its job. *}
nipkow@9932
   433
nipkow@16544
   434
subsection{*Finding Theorems\label{sec:find}*}
nipkow@16518
   435
nipkow@16523
   436
text{*\indexbold{finding theorems}\indexbold{searching theorems}
paulson@16560
   437
Isabelle's large database of proved theorems 
paulson@16560
   438
offers a powerful search engine. Its chief limitation is
nipkow@16518
   439
its restriction to the theories currently loaded.
nipkow@16518
   440
nipkow@16518
   441
\begin{pgnote}
nipkow@16523
   442
The search engine is started by clicking on Proof General's \pgmenu{Find} icon.
nipkow@16518
   443
You specify your search textually in the input buffer at the bottom
nipkow@16518
   444
of the window.
nipkow@16518
   445
\end{pgnote}
nipkow@16518
   446
paulson@16560
   447
The simplest form of search finds theorems containing specified
paulson@16560
   448
patterns.  A pattern can be any term (even
paulson@16560
   449
a single identifier).  It may contain ``\texttt{\_}'', a wildcard standing
paulson@16560
   450
for any term. Here are some
nipkow@16518
   451
examples:
nipkow@16518
   452
\begin{ttbox}
nipkow@16518
   453
length
nipkow@16518
   454
"_ # _ = _ # _"
nipkow@16518
   455
"_ + _"
nipkow@16518
   456
"_ * (_ - (_::nat))"
nipkow@16518
   457
\end{ttbox}
paulson@16560
   458
Specifying types, as shown in the last example, 
paulson@16560
   459
constrains searches involving overloaded operators.
nipkow@16518
   460
nipkow@16518
   461
\begin{warn}
nipkow@16518
   462
Always use ``\texttt{\_}'' rather than variable names: searching for
nipkow@16518
   463
\texttt{"x + y"} will usually not find any matching theorems
paulson@16560
   464
because they would need to contain \texttt{x} and~\texttt{y} literally.
paulson@16560
   465
When searching for infix operators, do not just type in the symbol,
paulson@16560
   466
such as~\texttt{+}, but a proper term such as \texttt{"_ + _"}.
paulson@16560
   467
This remark applies to more complicated syntaxes, too.
nipkow@16518
   468
\end{warn}
nipkow@16518
   469
paulson@16560
   470
If you are looking for rewrite rules (possibly conditional) that could
paulson@16560
   471
simplify some term, prefix the pattern with \texttt{simp:}.
nipkow@16518
   472
\begin{ttbox}
nipkow@16518
   473
simp: "_ * (_ + _)"
nipkow@16518
   474
\end{ttbox}
paulson@16560
   475
This finds \emph{all} equations---not just those with a \isa{simp} attribute---whose conclusion has the form
nipkow@16518
   476
@{text[display]"_ * (_ + _) = \<dots>"}
paulson@16560
   477
It only finds equations that can simplify the given pattern
paulson@16560
   478
at the root, not somewhere inside: for example, equations of the form
paulson@16560
   479
@{text"_ + _ = \<dots>"} do not match.
nipkow@16518
   480
paulson@16560
   481
You may also search for theorems by name---you merely
paulson@16560
   482
need to specify a substring. For example, you could search for all
nipkow@16518
   483
commutativity theorems like this:
nipkow@16518
   484
\begin{ttbox}
nipkow@16518
   485
name: comm
nipkow@16518
   486
\end{ttbox}
nipkow@16518
   487
This retrieves all theorems whose name contains \texttt{comm}.
nipkow@16518
   488
paulson@16560
   489
Search criteria can also be negated by prefixing them with ``\texttt{-}''.
paulson@16560
   490
For example,
nipkow@16518
   491
\begin{ttbox}
nipkow@16523
   492
-name: List
nipkow@16518
   493
\end{ttbox}
paulson@16560
   494
finds theorems whose name does not contain \texttt{List}. You can use this
paulson@16560
   495
to exclude particular theories from the search: the long name of
nipkow@16518
   496
a theorem contains the name of the theory it comes from.
nipkow@16518
   497
paulson@16560
   498
Finallly, different search criteria can be combined arbitrarily. 
webertj@20143
   499
The effect is conjuctive: Find returns the theorems that satisfy all of
paulson@16560
   500
the criteria. For example,
nipkow@16518
   501
\begin{ttbox}
nipkow@16518
   502
"_ + _"  -"_ - _"  -simp: "_ * (_ + _)"  name: assoc
nipkow@16518
   503
\end{ttbox}
paulson@16560
   504
looks for theorems containing plus but not minus, and which do not simplify
paulson@16560
   505
\mbox{@{text"_ * (_ + _)"}} at the root, and whose name contains \texttt{assoc}.
nipkow@16518
   506
nipkow@16544
   507
Further search criteria are explained in \S\ref{sec:find2}.
nipkow@16523
   508
nipkow@16523
   509
\begin{pgnote}
nipkow@16523
   510
Proof General keeps a history of all your search expressions.
nipkow@16523
   511
If you click on \pgmenu{Find}, you can use the arrow keys to scroll
nipkow@16523
   512
through previous searches and just modify them. This saves you having
nipkow@16523
   513
to type in lengthy expressions again and again.
nipkow@16523
   514
\end{pgnote}
nipkow@16518
   515
*}
nipkow@9932
   516
(*<*)
wenzelm@9922
   517
end
nipkow@9932
   518
(*>*)