Merged with new cookbook.pdf.
--- a/CookBook/FirstSteps.thy Wed Feb 25 10:13:05 2009 +0100
+++ b/CookBook/FirstSteps.thy Wed Feb 25 10:14:42 2009 +0100
@@ -890,10 +890,10 @@
text {*
Theorem attributes are @{text "[simp]"}, @{text "[OF \<dots>]"}, @{text
- "[symmetric]"} and so on. Such attributes are \emph{neither} tags \emph{or} flags
- annotated to theorems, but functions that do further processing once the
- theorems are proven. In particular, it is not possible to find out
- what are all theorems that have an given attribute in common, unless of course
+ "[symmetric]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags
+ annotated to theorems, but functions that do further processing once a
+ theorem is proven. In particular, it is not possible to find out
+ what are all theorems that have a given attribute in common, unless of course
the function behind the attribute stores the theorems in a retrivable
datastructure. This can be easily done by the recipe described in
\ref{rec:named}.
@@ -907,12 +907,23 @@
HOL.elim: declaration of Classical elimination rule
\<dots>"}
+ To explain how to write your own attribute, let us start with an extremely simple
+ version of the attribute @{text "[symmetric]"}. The purpose of this attribute is
+ to produce the ``symmetric'' version of an equation. The main function behind
+ this attribute is
*}
ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
text {*
- the setup
+ where the function @{ML "Thm.rule_attribute"} expects a function taking a
+ context (we ignore it in the code above) and a theorem (@{text thm}) and
+ returns another theorem (namely @{text thm} resolved with the rule
+ @{thm sym[no_vars]}). The function @{ML "Thm.rule_attribute"} then returns
+ an attribute (of type @{ML_type "attribute"}).
+
+ Before we can use the attribute, we need to set it up. This can be done
+ using the function @{ML Attrib.add_attributes} as follows.
*}
setup {*
@@ -920,16 +931,29 @@
[("my_sym", Attrib.no_args my_symmetric, "applying the sym rule")]
*}
-lemma test: "1 = Suc 0" by simp
+text {*
+ The attribute does not expect any further arguments (like @{text "[OF \<dots>]"} that
+ can take a list of theorems as argument). Therefore we use the function
+ @{ML Attrib.no_args}. Later on we will also consider attributes taking further
+ arguments. An example for the attribute @{text "[my_sym]"} is the proof
+*}
+
+lemma test[my_sym]: "2 = Suc (Suc 0)" by simp
text {*
- @{thm test[my_sym]}
+ which stores the theorem @{thm test} under the name @{thm [source] test}. We
+ can also use the attribute when referring to this theorem.
+
+ \begin{isabelle}
+ \isacommand{thm}~@{text "test[my_sym]"}\\
+ @{text "> "}~@{thm test[my_sym]}
+ \end{isabelle}
+
*}
text {*
What are:
-@{text "rule_attribute"}
@{text "declaration_attribute"}
--- a/CookBook/Intro.thy Wed Feb 25 10:13:05 2009 +0100
+++ b/CookBook/Intro.thy Wed Feb 25 10:14:42 2009 +0100
@@ -126,6 +126,7 @@
\item {\bf Sascha Böhme} contributed the recipes in \ref{rec:timeout},
\ref{rec:config}, \ref{rec:storingdata}, \ref{rec:external} and \ref{rec:oracle}.
+ He also wrote section \ref{sec:conversion}.
\item {\bf Jeremy Dawson} wrote the first version of the chapter
about parsing.
--- a/CookBook/Tactical.thy Wed Feb 25 10:13:05 2009 +0100
+++ b/CookBook/Tactical.thy Wed Feb 25 10:14:42 2009 +0100
@@ -1314,7 +1314,7 @@
one can say about them?)
*}
-section {* Conversions *}
+section {* Conversions\label{sec:conversion} *}
text {*
Conversions are meta-equalities depending on some input term. There type is
Binary file cookbook.pdf has changed