9993
|
1 |
%
|
|
2 |
\begin{isabellebody}%
|
|
3 |
\def\isabellecontext{simp}%
|
11866
|
4 |
\isamarkupfalse%
|
9993
|
5 |
%
|
10395
|
6 |
\isamarkupsection{Simplification%
|
|
7 |
}
|
11866
|
8 |
\isamarkuptrue%
|
9993
|
9 |
%
|
|
10 |
\begin{isamarkuptext}%
|
|
11 |
\label{sec:simplification-II}\index{simplification|(}
|
11494
|
12 |
This section describes features not covered until now. It also
|
|
13 |
outlines the simplification process itself, which can be helpful
|
|
14 |
when the simplifier does not do what you expect of it.%
|
9993
|
15 |
\end{isamarkuptext}%
|
11866
|
16 |
\isamarkuptrue%
|
9993
|
17 |
%
|
10878
|
18 |
\isamarkupsubsection{Advanced Features%
|
10395
|
19 |
}
|
11866
|
20 |
\isamarkuptrue%
|
9993
|
21 |
%
|
10878
|
22 |
\isamarkupsubsubsection{Congruence Rules%
|
10395
|
23 |
}
|
11866
|
24 |
\isamarkuptrue%
|
9993
|
25 |
%
|
|
26 |
\begin{isamarkuptext}%
|
|
27 |
\label{sec:simp-cong}
|
11494
|
28 |
While simplifying the conclusion $Q$
|
|
29 |
of $P \Imp Q$, it is legal use the assumption $P$.
|
|
30 |
For $\Imp$ this policy is hardwired, but
|
|
31 |
contextual information can also be made available for other
|
9993
|
32 |
operators. For example, \isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ xs\ {\isacharat}\ xs\ {\isacharequal}\ xs} simplifies to \isa{True} because we may use \isa{xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}} when simplifying \isa{xs\ {\isacharat}\ xs\ {\isacharequal}\ xs}. The generation of contextual information during simplification is
|
|
33 |
controlled by so-called \bfindex{congruence rules}. This is the one for
|
|
34 |
\isa{{\isasymlongrightarrow}}:
|
|
35 |
\begin{isabelle}%
|
10696
|
36 |
\ \ \ \ \ {\isasymlbrakk}P\ {\isacharequal}\ P{\isacharprime}{\isacharsemicolon}\ P{\isacharprime}\ {\isasymLongrightarrow}\ Q\ {\isacharequal}\ Q{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}P{\isacharprime}\ {\isasymlongrightarrow}\ Q{\isacharprime}{\isacharparenright}%
|
9993
|
37 |
\end{isabelle}
|
|
38 |
It should be read as follows:
|
|
39 |
In order to simplify \isa{P\ {\isasymlongrightarrow}\ Q} to \isa{P{\isacharprime}\ {\isasymlongrightarrow}\ Q{\isacharprime}},
|
|
40 |
simplify \isa{P} to \isa{P{\isacharprime}}
|
|
41 |
and assume \isa{P{\isacharprime}} when simplifying \isa{Q} to \isa{Q{\isacharprime}}.
|
|
42 |
|
|
43 |
Here are some more examples. The congruence rules for bounded
|
|
44 |
quantifiers supply contextual information about the bound variable:
|
|
45 |
\begin{isabelle}%
|
10696
|
46 |
\ \ \ \ \ {\isasymlbrakk}A\ {\isacharequal}\ B{\isacharsemicolon}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ B\ {\isasymLongrightarrow}\ P\ x\ {\isacharequal}\ Q\ x{\isasymrbrakk}\isanewline
|
10950
|
47 |
\isaindent{\ \ \ \ \ }{\isasymLongrightarrow}\ {\isacharparenleft}{\isasymforall}x{\isasymin}A{\isachardot}\ P\ x{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymforall}x{\isasymin}B{\isachardot}\ Q\ x{\isacharparenright}%
|
9993
|
48 |
\end{isabelle}
|
11494
|
49 |
One congruence rule for conditional expressions supplies contextual
|
11196
|
50 |
information for simplifying the \isa{then} and \isa{else} cases:
|
9993
|
51 |
\begin{isabelle}%
|
10696
|
52 |
\ \ \ \ \ {\isasymlbrakk}b\ {\isacharequal}\ c{\isacharsemicolon}\ c\ {\isasymLongrightarrow}\ x\ {\isacharequal}\ u{\isacharsemicolon}\ {\isasymnot}\ c\ {\isasymLongrightarrow}\ y\ {\isacharequal}\ v{\isasymrbrakk}\isanewline
|
10950
|
53 |
\isaindent{\ \ \ \ \ }{\isasymLongrightarrow}\ {\isacharparenleft}if\ b\ then\ x\ else\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ c\ then\ u\ else\ v{\isacharparenright}%
|
9993
|
54 |
\end{isabelle}
|
11494
|
55 |
An alternative congruence rule for conditional expressions
|
|
56 |
actually \emph{prevents} simplification of some arguments:
|
9993
|
57 |
\begin{isabelle}%
|
|
58 |
\ \ \ \ \ b\ {\isacharequal}\ c\ {\isasymLongrightarrow}\ {\isacharparenleft}if\ b\ then\ x\ else\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ c\ then\ x\ else\ y{\isacharparenright}%
|
|
59 |
\end{isabelle}
|
|
60 |
Only the first argument is simplified; the others remain unchanged.
|
|
61 |
This makes simplification much faster and is faithful to the evaluation
|
|
62 |
strategy in programming languages, which is why this is the default
|
11196
|
63 |
congruence rule for \isa{if}. Analogous rules control the evaluation of
|
9993
|
64 |
\isa{case} expressions.
|
|
65 |
|
11458
|
66 |
You can declare your own congruence rules with the attribute \attrdx{cong},
|
9993
|
67 |
either globally, in the usual manner,
|
|
68 |
\begin{quote}
|
|
69 |
\isacommand{declare} \textit{theorem-name} \isa{{\isacharbrackleft}cong{\isacharbrackright}}
|
|
70 |
\end{quote}
|
|
71 |
or locally in a \isa{simp} call by adding the modifier
|
|
72 |
\begin{quote}
|
|
73 |
\isa{cong{\isacharcolon}} \textit{list of theorem names}
|
|
74 |
\end{quote}
|
|
75 |
The effect is reversed by \isa{cong\ del} instead of \isa{cong}.
|
|
76 |
|
|
77 |
\begin{warn}
|
|
78 |
The congruence rule \isa{conj{\isacharunderscore}cong}
|
|
79 |
\begin{isabelle}%
|
10696
|
80 |
\ \ \ \ \ {\isasymlbrakk}P\ {\isacharequal}\ P{\isacharprime}{\isacharsemicolon}\ P{\isacharprime}\ {\isasymLongrightarrow}\ Q\ {\isacharequal}\ Q{\isacharprime}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}P\ {\isasymand}\ Q{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}P{\isacharprime}\ {\isasymand}\ Q{\isacharprime}{\isacharparenright}%
|
9993
|
81 |
\end{isabelle}
|
10878
|
82 |
\par\noindent
|
|
83 |
is occasionally useful but is not a default rule; you have to declare it explicitly.
|
9993
|
84 |
\end{warn}%
|
|
85 |
\end{isamarkuptext}%
|
11866
|
86 |
\isamarkuptrue%
|
9993
|
87 |
%
|
10878
|
88 |
\isamarkupsubsubsection{Permutative Rewrite Rules%
|
10395
|
89 |
}
|
11866
|
90 |
\isamarkuptrue%
|
9993
|
91 |
%
|
|
92 |
\begin{isamarkuptext}%
|
11494
|
93 |
\index{rewrite rules!permutative|bold}%
|
|
94 |
An equation is a \textbf{permutative rewrite rule} if the left-hand
|
9993
|
95 |
side and right-hand side are the same up to renaming of variables. The most
|
|
96 |
common permutative rule is commutativity: \isa{x\ {\isacharplus}\ y\ {\isacharequal}\ y\ {\isacharplus}\ x}. Other examples
|
|
97 |
include \isa{x\ {\isacharminus}\ y\ {\isacharminus}\ z\ {\isacharequal}\ x\ {\isacharminus}\ z\ {\isacharminus}\ y} in arithmetic and \isa{insert\ x\ {\isacharparenleft}insert\ y\ A{\isacharparenright}\ {\isacharequal}\ insert\ y\ {\isacharparenleft}insert\ x\ A{\isacharparenright}} for sets. Such rules are problematic because
|
|
98 |
once they apply, they can be used forever. The simplifier is aware of this
|
|
99 |
danger and treats permutative rules by means of a special strategy, called
|
|
100 |
\bfindex{ordered rewriting}: a permutative rewrite
|
10978
|
101 |
rule is only applied if the term becomes smaller with respect to a fixed
|
|
102 |
lexicographic ordering on terms. For example, commutativity rewrites
|
9993
|
103 |
\isa{b\ {\isacharplus}\ a} to \isa{a\ {\isacharplus}\ b}, but then stops because \isa{a\ {\isacharplus}\ b} is strictly
|
|
104 |
smaller than \isa{b\ {\isacharplus}\ a}. Permutative rewrite rules can be turned into
|
|
105 |
simplification rules in the usual manner via the \isa{simp} attribute; the
|
|
106 |
simplifier recognizes their special status automatically.
|
|
107 |
|
|
108 |
Permutative rewrite rules are most effective in the case of
|
10281
|
109 |
associative-commutative functions. (Associativity by itself is not
|
|
110 |
permutative.) When dealing with an AC-function~$f$, keep the
|
9993
|
111 |
following points in mind:
|
10281
|
112 |
\begin{itemize}\index{associative-commutative function}
|
9993
|
113 |
|
|
114 |
\item The associative law must always be oriented from left to right,
|
|
115 |
namely $f(f(x,y),z) = f(x,f(y,z))$. The opposite orientation, if
|
|
116 |
used with commutativity, can lead to nontermination.
|
|
117 |
|
|
118 |
\item To complete your set of rewrite rules, you must add not just
|
|
119 |
associativity~(A) and commutativity~(C) but also a derived rule, {\bf
|
|
120 |
left-com\-mut\-ativ\-ity} (LC): $f(x,f(y,z)) = f(y,f(x,z))$.
|
|
121 |
\end{itemize}
|
|
122 |
Ordered rewriting with the combination of A, C, and LC sorts a term
|
|
123 |
lexicographically:
|
|
124 |
\[\def\maps#1{~\stackrel{#1}{\leadsto}~}
|
|
125 |
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)) \]
|
|
126 |
|
|
127 |
Note that ordered rewriting for \isa{{\isacharplus}} and \isa{{\isacharasterisk}} on numbers is rarely
|
10878
|
128 |
necessary because the built-in arithmetic prover often succeeds without
|
11196
|
129 |
such tricks.%
|
9993
|
130 |
\end{isamarkuptext}%
|
11866
|
131 |
\isamarkuptrue%
|
9993
|
132 |
%
|
11494
|
133 |
\isamarkupsubsection{How the Simplifier Works%
|
10395
|
134 |
}
|
11866
|
135 |
\isamarkuptrue%
|
9993
|
136 |
%
|
|
137 |
\begin{isamarkuptext}%
|
|
138 |
\label{sec:SimpHow}
|
11494
|
139 |
Roughly speaking, the simplifier proceeds bottom-up: subterms are simplified
|
|
140 |
first. A conditional equation is only applied if its condition can be
|
|
141 |
proved, again by simplification. Below we explain some special features of
|
|
142 |
the rewriting process.%
|
9993
|
143 |
\end{isamarkuptext}%
|
11866
|
144 |
\isamarkuptrue%
|
9993
|
145 |
%
|
10878
|
146 |
\isamarkupsubsubsection{Higher-Order Patterns%
|
10395
|
147 |
}
|
11866
|
148 |
\isamarkuptrue%
|
9993
|
149 |
%
|
10186
|
150 |
\begin{isamarkuptext}%
|
|
151 |
\index{simplification rule|(}
|
|
152 |
So far we have pretended the simplifier can deal with arbitrary
|
11494
|
153 |
rewrite rules. This is not quite true. For reasons of feasibility,
|
|
154 |
the simplifier expects the
|
10186
|
155 |
left-hand side of each rule to be a so-called \emph{higher-order
|
11494
|
156 |
pattern}~\cite{nipkow-patterns}\indexbold{patterns!higher-order}.
|
|
157 |
This restricts where
|
10186
|
158 |
unknowns may occur. Higher-order patterns are terms in $\beta$-normal
|
11494
|
159 |
form. (This means there are no subterms of the form $(\lambda x. M)(N)$.)
|
|
160 |
Each occurrence of an unknown is of the form
|
10186
|
161 |
$\Var{f}~x@1~\dots~x@n$, where the $x@i$ are distinct bound
|
10978
|
162 |
variables. Thus all ordinary rewrite rules, where all unknowns are
|
11494
|
163 |
of base type, for example \isa{{\isacharquery}m\ {\isacharplus}\ {\isacharquery}n\ {\isacharplus}\ {\isacharquery}k\ {\isacharequal}\ {\isacharquery}m\ {\isacharplus}\ {\isacharparenleft}{\isacharquery}n\ {\isacharplus}\ {\isacharquery}k{\isacharparenright}}, are acceptable: if an unknown is
|
10186
|
164 |
of base type, it cannot have any arguments. Additionally, the rule
|
11494
|
165 |
\isa{{\isacharparenleft}{\isasymforall}x{\isachardot}\ {\isacharquery}P\ x\ {\isasymand}\ {\isacharquery}Q\ x{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isacharparenleft}{\isasymforall}x{\isachardot}\ {\isacharquery}P\ x{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}{\isasymforall}x{\isachardot}\ {\isacharquery}Q\ x{\isacharparenright}{\isacharparenright}} is also acceptable, in
|
10186
|
166 |
both directions: all arguments of the unknowns \isa{{\isacharquery}P} and
|
|
167 |
\isa{{\isacharquery}Q} are distinct bound variables.
|
|
168 |
|
11494
|
169 |
If the left-hand side is not a higher-order pattern, all is not lost.
|
|
170 |
The simplifier will still try to apply the rule provided it
|
|
171 |
matches directly: without much $\lambda$-calculus hocus
|
|
172 |
pocus. For example, \isa{{\isacharparenleft}{\isacharquery}f\ {\isacharquery}x\ {\isasymin}\ range\ {\isacharquery}f{\isacharparenright}\ {\isacharequal}\ True} rewrites
|
10186
|
173 |
\isa{g\ a\ {\isasymin}\ range\ g} to \isa{True}, but will fail to match
|
|
174 |
\isa{g{\isacharparenleft}h\ b{\isacharparenright}\ {\isasymin}\ range{\isacharparenleft}{\isasymlambda}x{\isachardot}\ g{\isacharparenleft}h\ x{\isacharparenright}{\isacharparenright}}. However, you can
|
11494
|
175 |
eliminate the offending subterms --- those that are not patterns ---
|
|
176 |
by adding new variables and conditions.
|
|
177 |
In our example, we eliminate \isa{{\isacharquery}f\ {\isacharquery}x} and obtain
|
|
178 |
\isa{{\isacharquery}y\ {\isacharequal}\ {\isacharquery}f\ {\isacharquery}x\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isacharquery}y\ {\isasymin}\ range\ {\isacharquery}f{\isacharparenright}\ {\isacharequal}\ True}, which is fine
|
10186
|
179 |
as a conditional rewrite rule since conditions can be arbitrary
|
11494
|
180 |
terms. However, this trick is not a panacea because the newly
|
11196
|
181 |
introduced conditions may be hard to solve.
|
10186
|
182 |
|
10878
|
183 |
There is no restriction on the form of the right-hand
|
10186
|
184 |
sides. They may not contain extraneous term or type variables, though.%
|
|
185 |
\end{isamarkuptext}%
|
11866
|
186 |
\isamarkuptrue%
|
9993
|
187 |
%
|
10878
|
188 |
\isamarkupsubsubsection{The Preprocessor%
|
10395
|
189 |
}
|
11866
|
190 |
\isamarkuptrue%
|
9993
|
191 |
%
|
|
192 |
\begin{isamarkuptext}%
|
10845
|
193 |
\label{sec:simp-preprocessor}
|
10878
|
194 |
When a theorem is declared a simplification rule, it need not be a
|
10186
|
195 |
conditional equation already. The simplifier will turn it into a set of
|
10878
|
196 |
conditional equations automatically. For example, \isa{f\ x\ {\isacharequal}\ g\ x\ {\isasymand}\ h\ x\ {\isacharequal}\ k\ x} becomes the two separate
|
|
197 |
simplification rules \isa{f\ x\ {\isacharequal}\ g\ x} and \isa{h\ x\ {\isacharequal}\ k\ x}. In
|
10186
|
198 |
general, the input theorem is converted as follows:
|
|
199 |
\begin{eqnarray}
|
10878
|
200 |
\neg P &\mapsto& P = \hbox{\isa{False}} \nonumber\\
|
10186
|
201 |
P \longrightarrow Q &\mapsto& P \Longrightarrow Q \nonumber\\
|
|
202 |
P \land Q &\mapsto& P,\ Q \nonumber\\
|
|
203 |
\forall x.~P~x &\mapsto& P~\Var{x}\nonumber\\
|
|
204 |
\forall x \in A.\ P~x &\mapsto& \Var{x} \in A \Longrightarrow P~\Var{x} \nonumber\\
|
|
205 |
\isa{if}\ P\ \isa{then}\ Q\ \isa{else}\ R &\mapsto&
|
|
206 |
P \Longrightarrow Q,\ \neg P \Longrightarrow R \nonumber
|
|
207 |
\end{eqnarray}
|
|
208 |
Once this conversion process is finished, all remaining non-equations
|
10878
|
209 |
$P$ are turned into trivial equations $P =\isa{True}$.
|
|
210 |
For example, the formula
|
|
211 |
\begin{center}\isa{{\isacharparenleft}p\ {\isasymlongrightarrow}\ t\ {\isacharequal}\ u\ {\isasymand}\ {\isasymnot}\ r{\isacharparenright}\ {\isasymand}\ s}\end{center}
|
10845
|
212 |
is converted into the three rules
|
10186
|
213 |
\begin{center}
|
10878
|
214 |
\isa{p\ {\isasymLongrightarrow}\ t\ {\isacharequal}\ u},\quad \isa{p\ {\isasymLongrightarrow}\ r\ {\isacharequal}\ False},\quad \isa{s\ {\isacharequal}\ True}.
|
10186
|
215 |
\end{center}
|
|
216 |
\index{simplification rule|)}
|
9993
|
217 |
\index{simplification|)}%
|
|
218 |
\end{isamarkuptext}%
|
11866
|
219 |
\isamarkuptrue%
|
|
220 |
\isamarkupfalse%
|
9993
|
221 |
\end{isabellebody}%
|
|
222 |
%%% Local Variables:
|
|
223 |
%%% mode: latex
|
|
224 |
%%% TeX-master: "root"
|
|
225 |
%%% End:
|