author | haftmann |
Thu, 20 May 2010 17:29:43 +0200 | |
changeset 37026 | 7e8979a155ae |
parent 31412 | f2e6b6526092 |
child 42814 | 5af15f1e2ef6 |
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 |
|
29650 | 9 |
uses "reify_data.ML" ("reflection.ML") |
20319 | 10 |
begin |
11 |
||
29650 | 12 |
setup {* Reify_Data.setup *} |
23546
c8a1bd9585a0
reflection and reification methods now manage Context data
chaieb
parents:
22199
diff
changeset
|
13 |
|
20374 | 14 |
lemma ext2: "(\<forall>x. f x = g x) \<Longrightarrow> f = g" |
15 |
by (blast intro: ext) |
|
22199
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
chaieb
parents:
21879
diff
changeset
|
16 |
|
20374 | 17 |
use "reflection.ML" |
18 |
||
30549 | 19 |
method_setup reify = {* |
20 |
Attrib.thms -- |
|
21 |
Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")")) >> |
|
22 |
(fn (eqs, to) => fn ctxt => SIMPLE_METHOD' (Reflection.genreify_tac ctxt (eqs @ (fst (Reify_Data.get ctxt))) to)) |
|
20319 | 23 |
*} "partial automatic reification" |
24 |
||
31412
f2e6b6526092
Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents:
30549
diff
changeset
|
25 |
method_setup reflection = {* |
f2e6b6526092
Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents:
30549
diff
changeset
|
26 |
let |
29650 | 27 |
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); |
28 |
val onlyN = "only"; |
|
29 |
val rulesN = "rules"; |
|
30 |
val any_keyword = keyword onlyN || keyword rulesN; |
|
31 |
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; |
|
32 |
val terms = thms >> map (term_of o Drule.dest_term); |
|
30549 | 33 |
in |
34 |
thms -- |
|
35 |
Scan.optional (keyword rulesN |-- thms) [] -- |
|
36 |
Scan.option (keyword onlyN |-- Args.term) >> |
|
37 |
(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
|
38 |
let |
f2e6b6526092
Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents:
30549
diff
changeset
|
39 |
val (ceqs,cths) = Reify_Data.get ctxt |
30549 | 40 |
val corr_thms = ths@cths |
41 |
val raw_eqs = eqs@ceqs |
|
42 |
in SIMPLE_METHOD' (Reflection.reflection_tac ctxt corr_thms raw_eqs to) end) |
|
43 |
end |
|
44 |
*} "reflection" |
|
29650 | 45 |
|
20319 | 46 |
end |