| author | wenzelm | 
| Fri, 14 Dec 2012 16:33:22 +0100 | |
| changeset 50530 | 6266e44b3396 | 
| parent 48891 | c0eafbd55de3 | 
| child 51542 | 738598beeb26 | 
| permissions | -rw-r--r-- | 
| 29650 | 1  | 
(* Title: HOL/Library/Reflection.thy  | 
| 20319 | 2  | 
Author: Amine Chaieb, TU Muenchen  | 
3  | 
*)  | 
|
4  | 
||
5  | 
header {* Generic reflection and reification *}
 | 
|
6  | 
||
7  | 
theory Reflection  | 
|
8  | 
imports Main  | 
|
9  | 
begin  | 
|
10  | 
||
| 48891 | 11  | 
ML_file "~~/src/HOL/Library/reify_data.ML"  | 
12  | 
ML_file "~~/src/HOL/Library/reflection.ML"  | 
|
13  | 
||
| 29650 | 14  | 
setup {* Reify_Data.setup *}
 | 
| 
23546
 
c8a1bd9585a0
reflection and reification methods now manage Context data
 
chaieb 
parents: 
22199 
diff
changeset
 | 
15  | 
|
| 30549 | 16  | 
method_setup reify = {*
 | 
17  | 
Attrib.thms --  | 
|
18  | 
    Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")")) >>
 | 
|
| 46765 | 19  | 
(fn (eqs, to) => fn ctxt => SIMPLE_METHOD' (Reflection.genreify_tac ctxt (eqs @ fst (Reify_Data.get ctxt)) to))  | 
| 20319 | 20  | 
*} "partial automatic reification"  | 
21  | 
||
| 
31412
 
f2e6b6526092
Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
 
hoelzl 
parents: 
30549 
diff
changeset
 | 
22  | 
method_setup reflection = {*
 | 
| 
 
f2e6b6526092
Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
 
hoelzl 
parents: 
30549 
diff
changeset
 | 
23  | 
let  | 
| 29650 | 24  | 
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();  | 
25  | 
val onlyN = "only";  | 
|
26  | 
val rulesN = "rules";  | 
|
27  | 
val any_keyword = keyword onlyN || keyword rulesN;  | 
|
28  | 
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;  | 
|
29  | 
val terms = thms >> map (term_of o Drule.dest_term);  | 
|
| 30549 | 30  | 
in  | 
31  | 
thms --  | 
|
32  | 
Scan.optional (keyword rulesN |-- thms) [] --  | 
|
33  | 
Scan.option (keyword onlyN |-- Args.term) >>  | 
|
| 46765 | 34  | 
(fn ((eqs, ths), to) => fn ctxt =>  | 
| 
31412
 
f2e6b6526092
Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
 
hoelzl 
parents: 
30549 
diff
changeset
 | 
35  | 
let  | 
| 46765 | 36  | 
val (ceqs, cths) = Reify_Data.get ctxt  | 
37  | 
val corr_thms = ths @ cths  | 
|
38  | 
val raw_eqs = eqs @ ceqs  | 
|
| 30549 | 39  | 
in SIMPLE_METHOD' (Reflection.reflection_tac ctxt corr_thms raw_eqs to) end)  | 
40  | 
end  | 
|
| 42814 | 41  | 
*}  | 
| 29650 | 42  | 
|
| 20319 | 43  | 
end  | 
| 46764 | 44  |