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