Theory Unification_Hints_Reification_Examples
section ‹Examples: Reification Via Unification Hints›
theory Unification_Hints_Reification_Examples
imports
HOL.Rat
ML_Unification_HOL_Setup
Unify_Fact_Tactic
Unify_Resolve_Tactics
begin
paragraph ‹Summary›
text ‹Reification via unification hints. For an introduction to unification hints refer
to \<^cite>‹"unif-hints"›. We support a generalisation of unification hints as described
in @{theory ML_Unification.ML_Unification_Hints}.›
subsection ‹Setup›
text ‹One-time setup to obtain a unifier with unification hints for the purpose of reification.›
ML‹
@{functor_instance struct_name = Reification_Unification_Hints
and functor_name = Term_Index_Unification_Hints
and id = ‹"reify"›
and more_args = ‹
structure TI = Discrimination_Tree
val init_args = {
concl_unifier = NONE, (*will be set later*)
prems_unifier = NONE, (*will be set later*)
normalisers = SOME Higher_Order_Pattern_Unification.norms_unify,
(*only retrieve hints based on hints' left-hand side*)
retrieval = SOME (Term_Index_Unification_Hints_Args.mk_retrieval_sym
(Term_Index_Unification_Hints_Args.retrieve_left TI.unifiables) TI.norm_term),
hint_preprocessor = SOME (Standard_Unification_Hints.get_hint_preprocessor
(Context.the_generic_context ()))
}›}
val reify_unify = Unification_Combinator.add_fallback_unifier
(fn unif_theory =>
Higher_Order_Pattern_Unification.e_unify Unification_Util.unify_types unif_theory unif_theory
|> Type_Unification.e_unify Unification_Util.unify_types)
(Reification_Unification_Hints.try_hints
|> Unification_Combinator.norm_unifier
(Unification_Util.inst_norm_term' Higher_Order_Pattern_Unification.norms_unify))
›
local_setup ‹Reification_Unification_Hints.setup_attribute NONE›
text ‹Premises of hints should again be unified by the reification unifier.›
declare [[reify_uhint where prems_unifier = reify_unify]]
subsection ‹Formulas with Quantifiers and Environment›
text ‹The following example is taken from HOL-Library.Reflection\_Examples. It is recommended
to compare the approach presented here with the reflection tactic presented in said theory.›