src/HOL/ex/Reflection.thy
author wenzelm
Thu, 05 Jul 2007 20:01:29 +0200
changeset 23592 ba0912262b2c
parent 23546 c8a1bd9585a0
child 23607 6a8fb529b542
permissions -rw-r--r--
renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac; simplified ObjectLogic.atomize;

(*
    ID:         $Id$
    Author:     Amine Chaieb, TU Muenchen
*)

header {* Generic reflection and reification *}

theory Reflection
imports Main
  uses "reflection_data.ML" ("reflection.ML")
begin

setup {* Reify_Data.setup*}


lemma ext2: "(\<forall>x. f x = g x) \<Longrightarrow> f = g"
  by (blast intro: ext)

use "reflection.ML"
ML{* Reify_Data.get @{context}*}

method_setup reify = {*
  fn src =>
    Method.syntax (Attrib.thms --
      Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")"))) src #>
  (fn ((eqs, to), ctxt) => Method.SIMPLE_METHOD' (Reflection.genreify_tac ctxt (eqs @ Reify_Data.get ctxt) to))
*} "partial automatic reification"

method_setup reflection = {*
  fn src =>
    Method.syntax (Attrib.thms --
      Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")"))) src #>
  (fn ((ths, to), ctxt) => Method.SIMPLE_METHOD' (Reflection.reflection_tac ctxt (ths @ Reify_Data.get ctxt) to))
*} "reflection method"

end