src/HOL/ex/reflection_data.ML
author wenzelm
Tue, 10 Jul 2007 23:29:43 +0200
changeset 23719 ccd9cb15c062
parent 23647 89286c4e7928
child 24045 bead02a55952
permissions -rw-r--r--
more markup for inner and outer syntax; added enclose;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     1
(*  Title:      HOL/ex/reflection_data.ML
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     2
    ID:         $Id$
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     3
    Author:     Amine Chaieb, TU Muenchen
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     4
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     5
Data for the reification and  reflection Methods
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     6
*)
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     7
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     8
signature REIFY_DATA =
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
     9
sig
23647
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    10
  type entry = (thm list) * (thm list)
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    11
  val get: Proof.context -> entry
23606
60950b22dcbf Cleaned add and del attributes
chaieb
parents: 23545
diff changeset
    12
  val del: attribute
60950b22dcbf Cleaned add and del attributes
chaieb
parents: 23545
diff changeset
    13
  val add: attribute 
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    14
  val setup: theory -> theory
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    15
end;
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    16
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    17
structure Reify_Data : REIFY_DATA =
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    18
struct
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    19
23647
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    20
type entry = (thm list) * (thm list);
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    21
fun appair f (x,y) = (f x, f y);
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    22
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    23
structure Data = GenericDataFun
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    24
( val name = "Reify-Data";
23647
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    25
  type T = entry
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    26
  val empty = ([],[]);
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    27
  val extend = I
23647
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    28
  fun merge _ = appair (Library.merge Thm.eq_thm));
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    29
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    30
val get = Data.get o Context.Proof;
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    31
23606
60950b22dcbf Cleaned add and del attributes
chaieb
parents: 23545
diff changeset
    32
val add = Thm.declaration_attribute (fn th => fn context => 
23647
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    33
  context |> Data.map (apfst (Drule.add_rule th )))
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    34
23606
60950b22dcbf Cleaned add and del attributes
chaieb
parents: 23545
diff changeset
    35
val del = Thm.declaration_attribute (fn th => fn context => 
23647
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    36
  context |> Data.map (apfst(Drule.del_rule th )))
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    37
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    38
val radd = Thm.declaration_attribute (fn th => fn context => 
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    39
  context |> Data.map (apsnd(Drule.add_rule th )))
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    40
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    41
val rdel = Thm.declaration_attribute (fn th => fn context => 
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    42
  context |> Data.map (apsnd(Drule.del_rule th )))
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    43
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    44
(* concrete syntax *)
23606
60950b22dcbf Cleaned add and del attributes
chaieb
parents: 23545
diff changeset
    45
(*
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    46
local
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    47
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    48
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    49
val constsN = "consts";
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    50
val addN = "add";
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    51
val delN = "del";
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    52
val any_keyword = keyword constsN || keyword addN || keyword delN;
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    53
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    54
val terms = thms >> map (term_of o Drule.dest_term);
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    55
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    56
fun optional scan = Scan.optional scan [];
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    57
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    58
in
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    59
fun att_syntax src = src |> Attrib.syntax
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    60
 ((Scan.lift (Args.$$$ "del") |-- thms) >> del ||
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    61
  optional (keyword addN |-- thms) >> add)
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    62
end;
23606
60950b22dcbf Cleaned add and del attributes
chaieb
parents: 23545
diff changeset
    63
*)
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    64
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    65
(* theory setup *)
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    66
 val setup =
23647
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    67
  Attrib.add_attributes [("reify", Attrib.add_del_args add del, "Reify data"),
89286c4e7928 Context data now consist of eqations for reification and equations for correctness theorems for reflection; Added attribute reflection to store correctness theorems to be trued by default
chaieb
parents: 23606
diff changeset
    68
                         ("reflection", Attrib.add_del_args radd rdel, "Reflection data")];
23545
2fddae6056ab Context Data for the reflection and reification methods
chaieb
parents:
diff changeset
    69
end;