| 104 |      1 | %% $Id$
 | 
|  |      2 | \chapter{Substitution Tactics} \label{substitution}
 | 
| 323 |      3 | \index{tactics!substitution|(}\index{equality|(}
 | 
|  |      4 | 
 | 
| 104 |      5 | Replacing equals by equals is a basic form of reasoning.  Isabelle supports
 | 
| 332 |      6 | several kinds of equality reasoning.  {\bf Substitution} means replacing
 | 
| 104 |      7 | free occurrences of~$t$ by~$u$ in a subgoal.  This is easily done, given an
 | 
| 332 |      8 | equality $t=u$, provided the logic possesses the appropriate rule.  The
 | 
|  |      9 | tactic {\tt hyp_subst_tac} performs substitution even in the assumptions.
 | 
|  |     10 | But it works via object-level implication, and therefore must be specially
 | 
|  |     11 | set up for each suitable object-logic.
 | 
| 104 |     12 | 
 | 
|  |     13 | Substitution should not be confused with object-level {\bf rewriting}.
 | 
|  |     14 | Given equalities of the form $t=u$, rewriting replaces instances of~$t$ by
 | 
|  |     15 | corresponding instances of~$u$, and continues until it reaches a normal
 | 
|  |     16 | form.  Substitution handles `one-off' replacements by particular
 | 
| 332 |     17 | equalities while rewriting handles general equations.
 | 
| 104 |     18 | Chapter~\ref{simp-chap} discusses Isabelle's rewriting tactics.
 | 
|  |     19 | 
 | 
|  |     20 | 
 | 
| 323 |     21 | \section{Substitution rules}
 | 
|  |     22 | \index{substitution!rules}\index{*subst theorem}
 | 
|  |     23 | Many logics include a substitution rule of the form
 | 
| 104 |     24 | $$ \List{\Var{a}=\Var{b}; \Var{P}(\Var{a})} \Imp 
 | 
|  |     25 |    \Var{P}(\Var{b})  \eqno(subst)$$
 | 
|  |     26 | In backward proof, this may seem difficult to use: the conclusion
 | 
|  |     27 | $\Var{P}(\Var{b})$ admits far too many unifiers.  But, if the theorem {\tt
 | 
|  |     28 | eqth} asserts $t=u$, then \hbox{\tt eqth RS subst} is the derived rule
 | 
|  |     29 | \[ \Var{P}(t) \Imp \Var{P}(u). \]
 | 
|  |     30 | Provided $u$ is not an unknown, resolution with this rule is
 | 
|  |     31 | well-behaved.\footnote{Unifying $\Var{P}(u)$ with a formula~$Q$
 | 
|  |     32 | expresses~$Q$ in terms of its dependence upon~$u$.  There are still $2^k$
 | 
|  |     33 | unifiers, if $Q$ has $k$ occurrences of~$u$, but Isabelle ensures that
 | 
|  |     34 | the first unifier includes all the occurrences.}  To replace $u$ by~$t$ in
 | 
|  |     35 | subgoal~$i$, use
 | 
|  |     36 | \begin{ttbox} 
 | 
| 332 |     37 | resolve_tac [eqth RS subst] \(i\){\it.}
 | 
| 104 |     38 | \end{ttbox}
 | 
|  |     39 | To replace $t$ by~$u$ in
 | 
|  |     40 | subgoal~$i$, use
 | 
|  |     41 | \begin{ttbox} 
 | 
| 332 |     42 | resolve_tac [eqth RS ssubst] \(i\){\it,}
 | 
| 104 |     43 | \end{ttbox}
 | 
| 323 |     44 | where \tdxbold{ssubst} is the `swapped' substitution rule
 | 
| 104 |     45 | $$ \List{\Var{a}=\Var{b}; \Var{P}(\Var{b})} \Imp 
 | 
|  |     46 |    \Var{P}(\Var{a}).  \eqno(ssubst)$$
 | 
| 323 |     47 | If \tdx{sym} denotes the symmetry rule
 | 
| 104 |     48 | \(\Var{a}=\Var{b}\Imp\Var{b}=\Var{a}\), then {\tt ssubst} is just
 | 
|  |     49 | \hbox{\tt sym RS subst}.  Many logics with equality include the rules {\tt
 | 
|  |     50 | subst} and {\tt ssubst}, as well as {\tt refl}, {\tt sym} and {\tt trans}
 | 
| 323 |     51 | (for the usual equality laws).  Examples include {\tt FOL} and {\tt HOL},
 | 
|  |     52 | but not {\tt CTT} (Constructive Type Theory).
 | 
| 104 |     53 | 
 | 
|  |     54 | Elim-resolution is well-behaved with assumptions of the form $t=u$.
 | 
|  |     55 | To replace $u$ by~$t$ or $t$ by~$u$ in subgoal~$i$, use
 | 
|  |     56 | \begin{ttbox} 
 | 
| 332 |     57 | eresolve_tac [subst] \(i\)    {\rm or}    eresolve_tac [ssubst] \(i\){\it.}
 | 
| 104 |     58 | \end{ttbox}
 | 
|  |     59 | 
 | 
|  |     60 | 
 | 
|  |     61 | \section{Substitution in the hypotheses}
 | 
| 323 |     62 | \index{assumptions!substitution in}
 | 
| 104 |     63 | Substitution rules, like other rules of natural deduction, do not affect
 | 
|  |     64 | the assumptions.  This can be inconvenient.  Consider proving the subgoal
 | 
|  |     65 | \[ \List{c=a; c=b} \Imp a=b. \]
 | 
| 323 |     66 | Calling {\tt eresolve_tac\ts[ssubst]\ts\(i\)} simply discards the
 | 
| 104 |     67 | assumption~$c=a$, since $c$ does not occur in~$a=b$.  Of course, we can
 | 
| 323 |     68 | work out a solution.  First apply {\tt eresolve_tac\ts[subst]\ts\(i\)},
 | 
| 104 |     69 | replacing~$a$ by~$c$:
 | 
|  |     70 | \[ \List{c=b} \Imp c=b \]
 | 
|  |     71 | Equality reasoning can be difficult, but this trivial proof requires
 | 
|  |     72 | nothing more sophisticated than substitution in the assumptions.
 | 
| 323 |     73 | Object-logics that include the rule~$(subst)$ provide tactics for this
 | 
| 104 |     74 | purpose:
 | 
|  |     75 | \begin{ttbox} 
 | 
| 323 |     76 | hyp_subst_tac       : int -> tactic
 | 
|  |     77 | bound_hyp_subst_tac : int -> tactic
 | 
| 104 |     78 | \end{ttbox}
 | 
| 323 |     79 | \begin{ttdescription}
 | 
| 104 |     80 | \item[\ttindexbold{hyp_subst_tac} {\it i}] 
 | 
| 323 |     81 |   selects an equality assumption of the form $t=u$ or $u=t$, where $t$ is a
 | 
|  |     82 |   free variable or parameter.  Deleting this assumption, it replaces $t$
 | 
|  |     83 |   by~$u$ throughout subgoal~$i$, including the other assumptions.
 | 
|  |     84 | 
 | 
|  |     85 | \item[\ttindexbold{bound_hyp_subst_tac} {\it i}] 
 | 
|  |     86 |   is similar but only substitutes for parameters (bound variables).
 | 
|  |     87 |   Uses for this are discussed below.
 | 
|  |     88 | \end{ttdescription}
 | 
| 104 |     89 | The term being replaced must be a free variable or parameter.  Substitution
 | 
|  |     90 | for constants is usually unhelpful, since they may appear in other
 | 
|  |     91 | theorems.  For instance, the best way to use the assumption $0=1$ is to
 | 
|  |     92 | contradict a theorem that states $0\not=1$, rather than to replace 0 by~1
 | 
|  |     93 | in the subgoal!
 | 
|  |     94 | 
 | 
| 332 |     95 | Substitution for free variables is also unhelpful if they appear in the
 | 
| 104 |     96 | premises of a rule being derived --- the substitution affects object-level
 | 
|  |     97 | assumptions, not meta-level assumptions.  For instance, replacing~$a$
 | 
| 323 |     98 | by~$b$ could make the premise~$P(a)$ worthless.  To avoid this problem, use
 | 
|  |     99 | {\tt bound_hyp_subst_tac}; alternatively, call \ttindex{cut_facts_tac} to
 | 
|  |    100 | insert the atomic premises as object-level assumptions.
 | 
| 104 |    101 | 
 | 
|  |    102 | 
 | 
|  |    103 | \section{Setting up {\tt hyp_subst_tac}} 
 | 
|  |    104 | Many Isabelle object-logics, such as {\tt FOL}, {\tt HOL} and their
 | 
|  |    105 | descendants, come with {\tt hyp_subst_tac} already defined.  A few others,
 | 
|  |    106 | such as {\tt CTT}, do not support this tactic because they lack the
 | 
|  |    107 | rule~$(subst)$.  When defining a new logic that includes a substitution
 | 
|  |    108 | rule and implication, you must set up {\tt hyp_subst_tac} yourself.  It
 | 
|  |    109 | is packaged as the \ML{} functor \ttindex{HypsubstFun}, which takes the
 | 
| 323 |    110 | argument signature~{\tt HYPSUBST_DATA}:
 | 
| 104 |    111 | \begin{ttbox} 
 | 
|  |    112 | signature HYPSUBST_DATA =
 | 
|  |    113 |   sig
 | 
|  |    114 |   val subst      : thm
 | 
|  |    115 |   val sym        : thm
 | 
|  |    116 |   val rev_cut_eq : thm
 | 
|  |    117 |   val imp_intr   : thm
 | 
|  |    118 |   val rev_mp     : thm
 | 
|  |    119 |   val dest_eq    : term -> term*term
 | 
|  |    120 |   end;
 | 
|  |    121 | \end{ttbox}
 | 
|  |    122 | Thus, the functor requires the following items:
 | 
| 323 |    123 | \begin{ttdescription}
 | 
|  |    124 | \item[\tdxbold{subst}] should be the substitution rule
 | 
| 104 |    125 | $\List{\Var{a}=\Var{b};\; \Var{P}(\Var{a})} \Imp \Var{P}(\Var{b})$.
 | 
|  |    126 | 
 | 
| 323 |    127 | \item[\tdxbold{sym}] should be the symmetry rule
 | 
| 104 |    128 | $\Var{a}=\Var{b}\Imp\Var{b}=\Var{a}$.
 | 
|  |    129 | 
 | 
| 323 |    130 | \item[\tdxbold{rev_cut_eq}] should have the form
 | 
| 104 |    131 | $\List{\Var{a}=\Var{b};\; \Var{a}=\Var{b}\Imp\Var{R}} \Imp \Var{R}$.
 | 
|  |    132 | 
 | 
| 323 |    133 | \item[\tdxbold{imp_intr}] should be the implies introduction
 | 
| 104 |    134 | rule $(\Var{P}\Imp\Var{Q})\Imp \Var{P}\imp\Var{Q}$.
 | 
|  |    135 | 
 | 
| 323 |    136 | \item[\tdxbold{rev_mp}] should be the `reversed' implies elimination
 | 
| 104 |    137 | rule $\List{\Var{P};  \;\Var{P}\imp\Var{Q}} \Imp \Var{Q}$.
 | 
|  |    138 | 
 | 
|  |    139 | \item[\ttindexbold{dest_eq}] should return the pair~$(t,u)$ when
 | 
|  |    140 | applied to the \ML{} term that represents~$t=u$.  For other terms, it
 | 
| 323 |    141 | should raise exception~\xdx{Match}.
 | 
|  |    142 | \end{ttdescription}
 | 
|  |    143 | The functor resides in file {\tt Provers/hypsubst.ML} in the Isabelle
 | 
| 104 |    144 | distribution directory.  It is not sensitive to the precise formalization
 | 
|  |    145 | of the object-logic.  It is not concerned with the names of the equality
 | 
|  |    146 | and implication symbols, or the types of formula and terms.  Coding the
 | 
|  |    147 | function {\tt dest_eq} requires knowledge of Isabelle's representation of
 | 
|  |    148 | terms.  For {\tt FOL} it is defined by
 | 
|  |    149 | \begin{ttbox} 
 | 
| 286 |    150 | fun dest_eq (Const("Trueprop",_) $ (Const("op =",_)$t$u)) = (t,u)
 | 
| 104 |    151 | \end{ttbox}
 | 
| 148 |    152 | Here {\tt Trueprop} is the coercion from type~$o$ to type~$prop$, while
 | 
| 104 |    153 | \hbox{\tt op =} is the internal name of the infix operator~{\tt=}.
 | 
|  |    154 | Pattern-matching expresses the function concisely, using wildcards~({\tt_})
 | 
| 332 |    155 | for the types.
 | 
| 104 |    156 | 
 | 
| 323 |    157 | Here is how {\tt hyp_subst_tac} works.  Given a subgoal of the form
 | 
|  |    158 | \[ \List{P@1; \cdots ; t=u; \cdots ; P@n} \Imp Q, \] it locates a suitable
 | 
|  |    159 | equality assumption and moves it to the last position using elim-resolution
 | 
|  |    160 | on {\tt rev_cut_eq} (possibly re-orienting it using~{\tt sym}):
 | 
| 104 |    161 | \[ \List{P@1; \cdots ; P@n; t=u} \Imp Q \]
 | 
| 323 |    162 | Using $n$ calls of {\tt eresolve_tac\ts[rev_mp]}, it creates the subgoal
 | 
| 104 |    163 | \[ \List{t=u} \Imp P@1\imp \cdots \imp P@n \imp Q \]
 | 
| 323 |    164 | By {\tt eresolve_tac\ts[ssubst]}, it replaces~$t$ by~$u$ throughout:
 | 
| 104 |    165 | \[ P'@1\imp \cdots \imp P'@n \imp Q' \]
 | 
| 323 |    166 | Finally, using $n$ calls of \hbox{\tt resolve_tac\ts[imp_intr]}, it restores
 | 
| 104 |    167 | $P'@1$, \ldots, $P'@n$ as assumptions:
 | 
|  |    168 | \[ \List{P'@n; \cdots ; P'@1} \Imp Q' \]
 | 
|  |    169 | 
 | 
| 323 |    170 | \index{equality|)}\index{tactics!substitution|)}
 |