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