src/Doc/Tutorial/Advanced/simp2.thy
author haftmann
Fri, 03 Jan 2014 22:04:44 +0100
changeset 54929 f1ded3cea58d
parent 48985 5386df44a037
child 57512 cc97b347b301
permissions -rw-r--r--
proper context for simplifier invocations in code generation stack
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
     1
(*<*)
48506
af1dabad14c0 avoid clash of Advanced/simp.thy vs. Misc/simp.thy;
wenzelm
parents: 19792
diff changeset
     2
theory simp2 imports Main begin
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
     3
(*>*)
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
     4
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
     5
section{*Simplification*}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
     6
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
     7
text{*\label{sec:simplification-II}\index{simplification|(}
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
     8
This section describes features not covered until now.  It also
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
     9
outlines the simplification process itself, which can be helpful
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    10
when the simplifier does not do what you expect of it.
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    11
*}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    12
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
    13
subsection{*Advanced Features*}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    14
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
    15
subsubsection{*Congruence Rules*}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    16
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    17
text{*\label{sec:simp-cong}
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    18
While simplifying the conclusion $Q$
13191
05a9929ee10e *** empty log message ***
nipkow
parents: 11494
diff changeset
    19
of $P \Imp Q$, it is legal to use the assumption $P$.
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    20
For $\Imp$ this policy is hardwired, but 
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    21
contextual information can also be made available for other
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    22
operators. For example, @{prop"xs = [] --> xs@xs = xs"} simplifies to @{term
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    23
True} because we may use @{prop"xs = []"} when simplifying @{prop"xs@xs =
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    24
xs"}. The generation of contextual information during simplification is
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    25
controlled by so-called \bfindex{congruence rules}. This is the one for
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    26
@{text"\<longrightarrow>"}:
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    27
@{thm[display]imp_cong[no_vars]}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    28
It should be read as follows:
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    29
In order to simplify @{prop"P-->Q"} to @{prop"P'-->Q'"},
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    30
simplify @{prop P} to @{prop P'}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    31
and assume @{prop"P'"} when simplifying @{prop Q} to @{prop"Q'"}.
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    32
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    33
Here are some more examples.  The congruence rules for bounded
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    34
quantifiers supply contextual information about the bound variable:
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    35
@{thm[display,eta_contract=false,margin=60]ball_cong[no_vars]}
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    36
One congruence rule for conditional expressions supplies contextual
11196
bb4ede27fcb7 *** empty log message ***
nipkow
parents: 10978
diff changeset
    37
information for simplifying the @{text then} and @{text else} cases:
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    38
@{thm[display]if_cong[no_vars]}
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    39
An alternative congruence rule for conditional expressions
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    40
actually \emph{prevents} simplification of some arguments:
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    41
@{thm[display]if_weak_cong[no_vars]}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    42
Only the first argument is simplified; the others remain unchanged.
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    43
This makes simplification much faster and is faithful to the evaluation
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    44
strategy in programming languages, which is why this is the default
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
    45
congruence rule for @{text "if"}. Analogous rules control the evaluation of
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    46
@{text case} expressions.
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    47
11458
09a6c44a48ea numerous stylistic changes and indexing
paulson
parents: 11216
diff changeset
    48
You can declare your own congruence rules with the attribute \attrdx{cong},
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    49
either globally, in the usual manner,
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    50
\begin{quote}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    51
\isacommand{declare} \textit{theorem-name} @{text"[cong]"}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    52
\end{quote}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    53
or locally in a @{text"simp"} call by adding the modifier
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    54
\begin{quote}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    55
@{text"cong:"} \textit{list of theorem names}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    56
\end{quote}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    57
The effect is reversed by @{text"cong del"} instead of @{text cong}.
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    58
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    59
\begin{warn}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    60
The congruence rule @{thm[source]conj_cong}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    61
@{thm[display]conj_cong[no_vars]}
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
    62
\par\noindent
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
    63
is occasionally useful but is not a default rule; you have to declare it explicitly.
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    64
\end{warn}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    65
*}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    66
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
    67
subsubsection{*Permutative Rewrite Rules*}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    68
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    69
text{*
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    70
\index{rewrite rules!permutative|bold}%
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
    71
An equation is a \textbf{permutative rewrite rule} if the left-hand
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    72
side and right-hand side are the same up to renaming of variables.  The most
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    73
common permutative rule is commutativity: @{prop"x+y = y+x"}.  Other examples
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    74
include @{prop"(x-y)-z = (x-z)-y"} in arithmetic and @{prop"insert x (insert
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    75
y A) = insert y (insert x A)"} for sets. Such rules are problematic because
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    76
once they apply, they can be used forever. The simplifier is aware of this
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    77
danger and treats permutative rules by means of a special strategy, called
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    78
\bfindex{ordered rewriting}: a permutative rewrite
10978
5eebea8f359f *** empty log message ***
nipkow
parents: 10885
diff changeset
    79
rule is only applied if the term becomes smaller with respect to a fixed
5eebea8f359f *** empty log message ***
nipkow
parents: 10885
diff changeset
    80
lexicographic ordering on terms. For example, commutativity rewrites
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    81
@{term"b+a"} to @{term"a+b"}, but then stops because @{term"a+b"} is strictly
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    82
smaller than @{term"b+a"}.  Permutative rewrite rules can be turned into
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    83
simplification rules in the usual manner via the @{text simp} attribute; the
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    84
simplifier recognizes their special status automatically.
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    85
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    86
Permutative rewrite rules are most effective in the case of
10281
9554ce1c2e54 *** empty log message ***
nipkow
parents: 10186
diff changeset
    87
associative-commutative functions.  (Associativity by itself is not
9554ce1c2e54 *** empty log message ***
nipkow
parents: 10186
diff changeset
    88
permutative.)  When dealing with an AC-function~$f$, keep the
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    89
following points in mind:
10281
9554ce1c2e54 *** empty log message ***
nipkow
parents: 10186
diff changeset
    90
\begin{itemize}\index{associative-commutative function}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    91
  
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    92
\item The associative law must always be oriented from left to right,
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    93
  namely $f(f(x,y),z) = f(x,f(y,z))$.  The opposite orientation, if
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    94
  used with commutativity, can lead to nontermination.
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    95
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    96
\item To complete your set of rewrite rules, you must add not just
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    97
  associativity~(A) and commutativity~(C) but also a derived rule, {\bf
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    98
    left-com\-mut\-ativ\-ity} (LC): $f(x,f(y,z)) = f(y,f(x,z))$.
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
    99
\end{itemize}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   100
Ordered rewriting with the combination of A, C, and LC sorts a term
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   101
lexicographically:
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   102
\[\def\maps#1{~\stackrel{#1}{\leadsto}~}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   103
 f(f(b,c),a) \maps{A} f(b,f(c,a)) \maps{C} f(b,f(a,c)) \maps{LC} f(a,f(b,c)) \]
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   104
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   105
Note that ordered rewriting for @{text"+"} and @{text"*"} on numbers is rarely
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   106
necessary because the built-in arithmetic prover often succeeds without
11196
bb4ede27fcb7 *** empty log message ***
nipkow
parents: 10978
diff changeset
   107
such tricks.
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   108
*}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   109
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   110
subsection{*How the Simplifier Works*}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   111
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   112
text{*\label{sec:SimpHow}
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   113
Roughly speaking, the simplifier proceeds bottom-up: subterms are simplified
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   114
first.  A conditional equation is only applied if its condition can be
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   115
proved, again by simplification.  Below we explain some special features of
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   116
the rewriting process. 
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   117
*}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   118
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   119
subsubsection{*Higher-Order Patterns*}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   120
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   121
text{*\index{simplification rule|(}
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   122
So far we have pretended the simplifier can deal with arbitrary
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   123
rewrite rules. This is not quite true.  For reasons of feasibility,
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   124
the simplifier expects the
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   125
left-hand side of each rule to be a so-called \emph{higher-order
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   126
pattern}~\cite{nipkow-patterns}\indexbold{patterns!higher-order}. 
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   127
This restricts where
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   128
unknowns may occur.  Higher-order patterns are terms in $\beta$-normal
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   129
form.  (This means there are no subterms of the form $(\lambda x. M)(N)$.)  
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   130
Each occurrence of an unknown is of the form
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   131
$\Var{f}~x@1~\dots~x@n$, where the $x@i$ are distinct bound
10978
5eebea8f359f *** empty log message ***
nipkow
parents: 10885
diff changeset
   132
variables. Thus all ordinary rewrite rules, where all unknowns are
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   133
of base type, for example @{thm add_assoc}, are acceptable: if an unknown is
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   134
of base type, it cannot have any arguments. Additionally, the rule
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   135
@{text"(\<forall>x. ?P x \<and> ?Q x) = ((\<forall>x. ?P x) \<and> (\<forall>x. ?Q x))"} is also acceptable, in
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   136
both directions: all arguments of the unknowns @{text"?P"} and
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   137
@{text"?Q"} are distinct bound variables.
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   138
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   139
If the left-hand side is not a higher-order pattern, all is not lost.
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   140
The simplifier will still try to apply the rule provided it
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   141
matches directly: without much $\lambda$-calculus hocus
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   142
pocus.  For example, @{text"(?f ?x \<in> range ?f) = True"} rewrites
15904
a6fb4ddc05c7 introduced @{const ...} antiquotation
haftmann
parents: 13191
diff changeset
   143
@{term"g a \<in> range g"} to @{const True}, but will fail to match
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   144
@{text"g(h b) \<in> range(\<lambda>x. g(h x))"}.  However, you can
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   145
eliminate the offending subterms --- those that are not patterns ---
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   146
by adding new variables and conditions.
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   147
In our example, we eliminate @{text"?f ?x"} and obtain
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   148
 @{text"?y =
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   149
?f ?x \<Longrightarrow> (?y \<in> range ?f) = True"}, which is fine
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   150
as a conditional rewrite rule since conditions can be arbitrary
11494
23a118849801 revisions and indexing
paulson
parents: 11458
diff changeset
   151
terms.  However, this trick is not a panacea because the newly
11196
bb4ede27fcb7 *** empty log message ***
nipkow
parents: 10978
diff changeset
   152
introduced conditions may be hard to solve.
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   153
  
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   154
There is no restriction on the form of the right-hand
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   155
sides.  They may not contain extraneous term or type variables, though.
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   156
*}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   157
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   158
subsubsection{*The Preprocessor*}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   159
10845
3696bc935bbd *** empty log message ***
nipkow
parents: 10795
diff changeset
   160
text{*\label{sec:simp-preprocessor}
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   161
When a theorem is declared a simplification rule, it need not be a
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   162
conditional equation already.  The simplifier will turn it into a set of
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   163
conditional equations automatically.  For example, @{prop"f x =
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   164
g x & h x = k x"} becomes the two separate
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   165
simplification rules @{prop"f x = g x"} and @{prop"h x = k x"}. In
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   166
general, the input theorem is converted as follows:
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   167
\begin{eqnarray}
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   168
\neg P &\mapsto& P = \hbox{\isa{False}} \nonumber\\
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   169
P \longrightarrow Q &\mapsto& P \Longrightarrow Q \nonumber\\
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   170
P \land Q &\mapsto& P,\ Q \nonumber\\
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   171
\forall x.~P~x &\mapsto& P~\Var{x}\nonumber\\
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   172
\forall x \in A.\ P~x &\mapsto& \Var{x} \in A \Longrightarrow P~\Var{x} \nonumber\\
19792
e8e3da6d3ff7 quoted "if";
wenzelm
parents: 16417
diff changeset
   173
@{text "if"}\ P\ @{text then}\ Q\ @{text else}\ R &\mapsto&
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   174
 P \Longrightarrow Q,\ \neg P \Longrightarrow R \nonumber
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   175
\end{eqnarray}
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   176
Once this conversion process is finished, all remaining non-equations
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   177
$P$ are turned into trivial equations $P =\isa{True}$.
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   178
For example, the formula 
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   179
\begin{center}@{prop"(p \<longrightarrow> t=u \<and> ~r) \<and> s"}\end{center}
10845
3696bc935bbd *** empty log message ***
nipkow
parents: 10795
diff changeset
   180
is converted into the three rules
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   181
\begin{center}
10885
90695f46440b lcp's pass over the book, chapters 1-8
paulson
parents: 10845
diff changeset
   182
@{prop"p \<Longrightarrow> t = u"},\quad  @{prop"p \<Longrightarrow> r = False"},\quad  @{prop"s = True"}.
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   183
\end{center}
499637e8f2c6 *** empty log message ***
nipkow
parents: 9958
diff changeset
   184
\index{simplification rule|)}
9958
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   185
\index{simplification|)}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   186
*}
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   187
(*<*)
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   188
end
67f2920862c7 *** empty log message ***
nipkow
parents:
diff changeset
   189
(*>*)