--- a/doc-src/IsarImplementation/Thy/Tactic.thy Sun Jan 29 22:00:10 2012 +0100
+++ b/doc-src/IsarImplementation/Thy/Tactic.thy Sun Jan 29 22:12:25 2012 +0100
@@ -282,7 +282,11 @@
premises.
\item @{ML eresolve_tac}~@{text "thms i"} performs elim-resolution
- with the given theorems, which should normally be elimination rules.
+ with the given theorems, which are normally be elimination rules.
+
+ Note that @{ML "eresolve_tac [asm_rl]"} is equivalent to @{ML
+ assume_tac}, which facilitates mixing of assumption steps with
+ genuine eliminations.
\item @{ML dresolve_tac}~@{text "thms i"} performs
destruct-resolution with the given theorems, which should normally
--- a/doc-src/IsarImplementation/Thy/document/Tactic.tex Sun Jan 29 22:00:10 2012 +0100
+++ b/doc-src/IsarImplementation/Thy/document/Tactic.tex Sun Jan 29 22:12:25 2012 +0100
@@ -338,7 +338,10 @@
premises.
\item \verb|eresolve_tac|~\isa{thms\ i} performs elim-resolution
- with the given theorems, which should normally be elimination rules.
+ with the given theorems, which are normally be elimination rules.
+
+ Note that \verb|eresolve_tac [asm_rl]| is equivalent to \verb|assume_tac|, which facilitates mixing of assumption steps with
+ genuine eliminations.
\item \verb|dresolve_tac|~\isa{thms\ i} performs
destruct-resolution with the given theorems, which should normally
--- a/doc-src/Ref/tactic.tex Sun Jan 29 22:00:10 2012 +0100
+++ b/doc-src/Ref/tactic.tex Sun Jan 29 22:12:25 2012 +0100
@@ -66,23 +66,6 @@
(see Chapter~\ref{chap:simplification}).
\end{warn}
-\subsection{Theorems useful with tactics}
-\index{theorems!of pure theory}
-\begin{ttbox}
-asm_rl: thm
-\end{ttbox}
-\begin{ttdescription}
-\item[\tdx{asm_rl}]
-is $\psi\Imp\psi$. Under elim-resolution it does proof by assumption, and
-\hbox{\tt eresolve_tac (asm_rl::{\it thms}) {\it i}} is equivalent to
-\begin{ttbox}
-assume_tac {\it i} ORELSE eresolve_tac {\it thms} {\it i}
-\end{ttbox}
-
-\end{ttdescription}
-
-
-\section{Obscure tactics}
\subsection{Composition: resolution without lifting}
\index{tactics!for composition}