Theory Reflection_Examples
section ‹Examples for generic reflection and reification›
theory Reflection_Examples
imports Complex_Main "HOL-Library.Reflection"
begin
text ‹This theory presents two methods: reify and reflection›
text ‹
Consider an HOL type ‹σ›, the structure of which is not recongnisable
on the theory level. This is the case of \<^typ>‹bool›, arithmetical terms such as \<^typ>‹int›,
\<^typ>‹real› etc \dots In order to implement a simplification on terms of type ‹σ› we
often need its structure. Traditionnaly such simplifications are written in ML,
proofs are synthesized.
An other strategy is to declare an HOL datatype ‹τ› and an HOL function (the
interpretation) that maps elements of ‹τ› to elements of ‹σ›.
The functionality of ‹reify› then is, given a term ‹t› of type ‹σ›,
to compute a term ‹s› of type ‹τ›. For this it needs equations for the
interpretation.
N.B: All the interpretations supported by ‹reify› must have the type
‹'a list ⇒ τ ⇒ σ›. The method ‹reify› can also be told which subterm
of the current subgoal should be reified. The general call for ‹reify› is
‹reify eqs (t)›, where ‹eqs› are the defining equations of the interpretation
and ‹(t)› is an optional parameter which specifies the subterm to which reification
should be applied to. If ‹(t)› is abscent, ‹reify› tries to reify the whole
subgoal.
The method ‹reflection› uses ‹reify› and has a very similar signature:
‹reflection corr_thm eqs (t)›. Here again ‹eqs› and ‹(t)›
are as described above and ‹corr_thm› is a theorem proving
\<^prop>‹I vs (f t) = I vs t›. We assume that ‹I› is the interpretation
and ‹f› is some useful and executable simplification of type ‹τ ⇒ τ›.
The method ‹reflection› applies reification and hence the theorem \<^prop>‹t = I xs s›
and hence using ‹corr_thm› derives \<^prop>‹t = I xs (f s)›. It then uses
normalization by equational rewriting to prove \<^prop>‹f s = s'› which almost finishes
the proof of \<^prop>‹t = t'› where \<^prop>‹I xs s' = t'›.
›
text ‹Example 1 : Propositional formulae and NNF.›
text ‹The type ‹fm› represents simple propositional formulae:›