| 20319 |      1 | (*
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Amine Chaieb, TU Muenchen
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header {* Generic reflection and reification *}
 | 
|  |      7 | 
 | 
|  |      8 | theory Reflection
 | 
|  |      9 | imports Main
 | 
| 20374 |     10 | uses ("reflection.ML")
 | 
| 20319 |     11 | begin
 | 
|  |     12 | 
 | 
| 20374 |     13 | lemma ext2: "(\<forall>x. f x = g x) \<Longrightarrow> f = g"
 | 
|  |     14 |   by (blast intro: ext)
 | 
|  |     15 | use "reflection.ML"
 | 
|  |     16 | 
 | 
| 20319 |     17 | method_setup reify = {*
 | 
|  |     18 |   fn src =>
 | 
|  |     19 |     Method.syntax (Attrib.thms --
 | 
|  |     20 |       Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")"))) src #>
 | 
| 21879 |     21 |   (fn ((eqs, to), ctxt) => Method.SIMPLE_METHOD' (Reflection.genreify_tac ctxt eqs to))
 | 
| 20319 |     22 | *} "partial automatic reification"
 | 
|  |     23 | 
 | 
|  |     24 | method_setup reflection = {*
 | 
|  |     25 |   fn src =>
 | 
|  |     26 |     Method.syntax (Attrib.thms --
 | 
|  |     27 |       Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")"))) src #>
 | 
| 21879 |     28 |   (fn ((ths, to), ctxt) => Method.SIMPLE_METHOD' (Reflection.reflection_tac ctxt ths to))
 | 
| 20319 |     29 | *} "reflection method"
 | 
|  |     30 | 
 | 
|  |     31 | end
 |