diff -r 2e48182cc82c -r cf58b5b794b2 src/HOL/ex/Reflection_Examples.thy --- a/src/HOL/ex/Reflection_Examples.thy Sat Dec 26 15:44:14 2015 +0100 +++ b/src/HOL/ex/Reflection_Examples.thy Sat Dec 26 15:59:27 2015 +0100 @@ -11,40 +11,40 @@ text \This theory presents two methods: reify and reflection\ text \ -Consider an HOL type @{text \}, the structure of which is not recongnisable +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 @{text \} we +@{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 @{text \} and an HOL function (the -interpretation) that maps elements of @{text \} to elements of @{text \}. +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 @{text reify} then is, given a term @{text t} of type @{text \}, -to compute a term @{text s} of type @{text \}. For this it needs equations for the +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 @{text reify} must have the type -@{text "'a list \ \ \ \"}. The method @{text reify} can also be told which subterm -of the current subgoal should be reified. The general call for @{text reify} is -@{text "reify eqs (t)"}, where @{text eqs} are the defining equations of the interpretation -and @{text "(t)"} is an optional parameter which specifies the subterm to which reification -should be applied to. If @{text "(t)"} is abscent, @{text reify} tries to reify the whole +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 @{text reflection} uses @{text reify} and has a very similar signature: -@{text "reflection corr_thm eqs (t)"}. Here again @{text eqs} and @{text "(t)"} -are as described above and @{text corr_thm} is a theorem proving -@{prop "I vs (f t) = I vs t"}. We assume that @{text I} is the interpretation -and @{text f} is some useful and executable simplification of type @{text "\ \ \"}. -The method @{text reflection} applies reification and hence the theorem @{prop "t = I xs s"} -and hence using @{text corr_thm} derives @{prop "t = I xs (f s)"}. It then uses +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 @{text fm} represents simple propositional formulae:\ +text \The type \fm\ represents simple propositional formulae:\ datatype form = TrueF | FalseF | Less nat nat | And form form | Or form form | Neg form | ExQ form @@ -81,8 +81,8 @@ apply (reify Ifm.simps) oops -text \Method @{text reify} maps a @{typ bool} to an @{typ fm}. For this it needs the -semantics of @{text fm}, i.e.\ the rewrite rules in @{text Ifm.simps}.\ +text \Method \reify\ maps a @{typ bool} to an @{typ fm}. For this it needs the +semantics of \fm\, i.e.\ the rewrite rules in \Ifm.simps\.\ text \You can also just pick up a subterm to reify.\ lemma "Q \ (D \ F \ ((\ D) \ (\ F)))" @@ -134,7 +134,7 @@ text \Example 2: Simple arithmetic formulae\ -text \The type @{text num} reflects linear expressions over natural number\ +text \The type \num\ reflects linear expressions over natural number\ datatype num = C nat | Add num num | Mul nat num | Var nat | CN nat nat num text \This is just technical to make recursive definitions easier.\ @@ -161,15 +161,15 @@ lemma "4 * (2 * x + (y::nat)) + f a \ 0" apply (reify Inum.simps ("4 * (2 * x + (y::nat)) + f a")) oops -text \We're in a bad situation! @{text x}, @{text y} and @{text f} have been recongnized +text \We're in a bad situation! \x\, \y\ and \f\ have been recongnized as constants, which is correct but does not correspond to our intuition of the constructor C. It should encapsulate constants, i.e. numbers, i.e. numerals.\ -text \So let's leave the @{text "Inum_C"} equation at the end and see what happens \dots\ +text \So let's leave the \Inum_C\ equation at the end and see what happens \dots\ lemma "4 * (2 * x + (y::nat)) \ 0" apply (reify Inum_Var Inum_Add Inum_Mul Inum_CN Inum_C ("4 * (2 * x + (y::nat))")) oops -text \Hm, let's specialize @{text Inum_C} with numerals.\ +text \Hm, let's specialize \Inum_C\ with numerals.\ lemma Inum_number: "Inum (C (numeral t)) vs = numeral t" by simp lemmas Inum_eqs = Inum_Var Inum_Add Inum_Mul Inum_CN Inum_number @@ -185,7 +185,7 @@ apply (reify Inum_eqs ("1 * (2 * x + (y::nat) + 0 + 1)")) oops -text \Oh!! 0 is not a variable \dots\ Oh! 0 is not a @{text "numeral"} \dots\ thing. +text \Oh!! 0 is not a variable \dots\ Oh! 0 is not a \numeral\ \dots\ thing. The same for 1. So let's add those equations, too.\ lemma Inum_01: "Inum (C 0) vs = 0" "Inum (C 1) vs = 1" "Inum (C(Suc n)) vs = Suc n" @@ -200,7 +200,7 @@ oops text \Okay, let's try reflection. Some simplifications on @{typ num} follow. You can - skim until the main theorem @{text linum}.\ + skim until the main theorem \linum\.\ fun lin_add :: "num \ num \ num" where