--- a/CookBook/Tactical.thy Thu Mar 12 18:39:10 2009 +0000
+++ b/CookBook/Tactical.thy Fri Mar 13 01:15:55 2009 +0100
@@ -1762,9 +1762,9 @@
Note that the actual response in this example is @{term "2 + 10 \<equiv> 2 + (10::nat)"},
since the pretty-printer for @{ML_type cterm}s already beta-normalises terms.
- But by the way how we constructed the term (using the function
- @{ML Thm.capply}, which is the application @{ML $} for @{ML_type cterm}s),
- we can be sure the left-hand side must contain beta-redexes. Indeed
+ But how we constructed the term (using the function
+ @{ML Thm.capply}, which is the application @{ML $} for @{ML_type cterm}s)
+ ensures that the left-hand side must contain beta-redexes. Indeed
if we obtain the ``raw'' representation of the produced theorem, we
can see the difference:
@@ -1808,7 +1808,7 @@
text {*
You can see how this function works in the example rewriting
- the @{ML_type cterm} @{term "True \<and> (Foo \<longrightarrow> Bar)"} to @{term "Foo \<longrightarrow> Bar"}.
+ @{term "True \<and> (Foo \<longrightarrow> Bar)"} to @{term "Foo \<longrightarrow> Bar"}.
@{ML_response_fake [display,gray]
"let
@@ -1861,6 +1861,14 @@
does not fail, however, because the combinator @{ML Conv.else_conv} will then
try out @{ML Conv.all_conv}, which always succeeds.
+ The conversion combinator @{ML Conv.try_conv} constructs a conversion
+ which is tried out on a term, but in case of failure just does nothing.
+ For example
+
+ @{ML_response_fake [display,gray]
+ "Conv.try_conv (Conv.rewr_conv @{thm true_conj1}) @{cterm \"True \<or> P\"}"
+ "True \<or> P \<equiv> True \<or> P"}
+
Apart from the function @{ML beta_conversion in Thm}, which is able to fully
beta-normalise a term, the conversions so far are restricted in that they
only apply to the outer-most level of a @{ML_type cterm}. In what follows we
@@ -2028,8 +2036,6 @@
and simprocs; the advantage of conversions, however, is that you never have
to worry about non-termination.
- (FIXME: explain @{ML Conv.try_conv})
-
\begin{exercise}\label{ex:addconversion}
Write a tactic that does the same as the simproc in exercise
\ref{ex:addsimproc}, but is based in conversions. That means replace terms
@@ -2038,7 +2044,7 @@
\end{exercise}
\begin{exercise}\label{ex:compare}
- Compare your solutions of Exercises~\ref{addsimproc} and \ref{ex:addconversion},
+ Compare your solutions of Exercises~\ref{ex:addsimproc} and \ref{ex:addconversion},
and try to determine which way of rewriting such terms is faster. For this you might
have to construct quite large terms. Also see Recipe \ref{rec:timing} for information
about timing.