author | wenzelm |
Tue, 23 Dec 2008 00:56:03 +0100 | |
changeset 29152 | 89b0803404d7 |
parent 23649 | 4d865f3e4405 |
permissions | -rw-r--r-- |
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 |
|
23546
c8a1bd9585a0
reflection and reification methods now manage Context data
chaieb
parents:
22199
diff
changeset
|
10 |
uses "reflection_data.ML" ("reflection.ML") |
20319 | 11 |
begin |
12 |
||
23546
c8a1bd9585a0
reflection and reification methods now manage Context data
chaieb
parents:
22199
diff
changeset
|
13 |
setup {* Reify_Data.setup*} |
c8a1bd9585a0
reflection and reification methods now manage Context data
chaieb
parents:
22199
diff
changeset
|
14 |
|
c8a1bd9585a0
reflection and reification methods now manage Context data
chaieb
parents:
22199
diff
changeset
|
15 |
|
20374 | 16 |
lemma ext2: "(\<forall>x. f x = g x) \<Longrightarrow> f = g" |
17 |
by (blast intro: ext) |
|
22199
b617ddd200eb
Now deals with simples cases where the input equations contain type variables
chaieb
parents:
21879
diff
changeset
|
18 |
|
20374 | 19 |
use "reflection.ML" |
20 |
||
20319 | 21 |
method_setup reify = {* |
22 |
fn src => |
|
23 |
Method.syntax (Attrib.thms -- |
|
23649
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
24 |
Scan.option (Scan.lift (Args.$$$ "(") |-- Args.term --| Scan.lift (Args.$$$ ")") )) src #> |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
25 |
(fn ((eqs, to), ctxt) => Method.SIMPLE_METHOD' (Reflection.genreify_tac ctxt (eqs @ (fst (Reify_Data.get ctxt))) to)) |
20319 | 26 |
*} "partial automatic reification" |
27 |
||
23649
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
28 |
method_setup reflection = {* |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
29 |
let |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
30 |
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
31 |
val onlyN = "only"; |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
32 |
val rulesN = "rules"; |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
33 |
val any_keyword = keyword onlyN || keyword rulesN; |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
34 |
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
35 |
val terms = thms >> map (term_of o Drule.dest_term); |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
36 |
fun optional scan = Scan.optional scan []; |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
37 |
in |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
38 |
fn src => |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
39 |
Method.syntax (thms -- optional (keyword rulesN |-- thms) -- Scan.option (keyword onlyN |-- Args.term)) src #> |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
40 |
(fn (((eqs,ths),to), ctxt) => |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
41 |
let |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
42 |
val (ceqs,cths) = Reify_Data.get ctxt |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
43 |
val corr_thms = ths@cths |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
44 |
val raw_eqs = eqs@ceqs |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
45 |
in Method.SIMPLE_METHOD' (Reflection.reflection_tac ctxt corr_thms raw_eqs to) |
4d865f3e4405
Changed syntax for the reflection method (reify unchanged). reflection takes an optional set of theorems for reification then several correctness theorems are specified by rules:'' then the optional term is specified by only:
chaieb
parents:
23607
diff
changeset
|
46 |
end) end |
20319 | 47 |
*} "reflection method" |
48 |
end |