Theorem tags and attributes.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/attribute.ML Fri Apr 03 14:35:39 1998 +0200
@@ -0,0 +1,160 @@
+(* Title: Pure/attribute.ML
+ ID: $Id$
+ Author: Markus Wenzel, TU Muenchen
+
+Theorem tags and attributes.
+*)
+
+signature BASIC_ATTRIBUTE =
+sig
+ type tag
+ type tthm
+ type 'a attribute
+ val print_attributes: theory -> unit
+end;
+
+signature ATTRIBUTE =
+sig
+ include BASIC_ATTRIBUTE
+ val thm_of: tthm -> thm
+ val tthm_of: thm -> tthm
+ val apply: ('a * tthm) * 'a attribute list -> ('a * tthm)
+ val pretty_tthm: tthm -> Pretty.T
+ val tag: tag -> 'a attribute
+ val simple: string -> 'a attribute
+ val lemma: tag
+ val assumption: tag
+ val tag_lemma: 'a attribute
+ val tag_assumption: 'a attribute
+ val setup: theory -> theory
+ val global_attr: theory -> xstring -> string list -> theory attribute
+ val local_attr: theory -> xstring -> string list -> object Symtab.table attribute
+ val add_attrs: (bstring * ((string list -> theory attribute) *
+ (string list -> object Symtab.table attribute))) list -> theory -> theory
+end;
+
+structure Attribute: ATTRIBUTE =
+struct
+
+
+(** tags and attributes **)
+
+type tag = string * string list;
+type tthm = thm * tag list;
+type 'a attribute = 'a * tthm -> 'a * tthm;
+
+fun thm_of (thm, _) = thm;
+fun tthm_of thm = (thm, []);
+
+fun apply x_attrs = foldl (op |>) x_attrs;
+
+
+(* display tagged theorems *)
+
+fun pretty_tag (name, []) = Pretty.str name
+ | pretty_tag (name, args) = Pretty.block
+ [Pretty.str name, Pretty.list "(" ")" (map Pretty.str args)];
+
+val pretty_tags = Pretty.list "[" "]" o map pretty_tag;
+
+fun pretty_tthm (thm, []) = Pretty.quote (Display.pretty_thm thm)
+ | pretty_tthm (thm, tags) = Pretty.block
+ [Pretty.quote (Display.pretty_thm thm), Pretty.brk 1, pretty_tags tags];
+
+
+(* basic attributes *)
+
+fun tag name_args (x, (thm, tags)) = (x, (thm, name_args ins tags));
+fun simple name = tag (name, []);
+
+val lemma = ("lemma", []);
+val assumption = ("assumption", []);
+fun tag_lemma x = tag lemma x;
+fun tag_assumption x = tag assumption x;
+
+
+
+(** theory data **)
+
+(* data kind 'attributes' *)
+
+val attributesK = "Pure/attributes";
+
+exception Attributes of
+ {space: NameSpace.T,
+ attrs:
+ ((string list -> theory attribute) *
+ (string list -> object Symtab.table attribute)) Symtab.table};
+
+fun print_attributes thy = Display.print_data thy attributesK;
+
+
+(* setup *)
+
+local
+ val empty = Attributes {space = NameSpace.empty, attrs = Symtab.empty};
+
+ fun prep_ext (x as Attributes _) = x;
+
+ fun merge (Attributes {space = space1, attrs = attrs1},
+ Attributes {space = space2, attrs = attrs2}) =
+ Attributes {
+ space = NameSpace.merge (space1, space2),
+ attrs = Symtab.merge (K true) (attrs1, attrs2)};
+
+ fun print _ (Attributes {space, attrs}) =
+ (Pretty.writeln (Display.pretty_name_space ("attribute name space", space));
+ Pretty.writeln (Pretty.strs ("attributes:" :: map fst (Symtab.dest attrs))));
+in
+ val setup = Theory.init_data [(attributesK, (empty, prep_ext, merge, print))];
+end;
+
+
+(* get data record *)
+
+fun get_attributes_sg sg =
+ (case Sign.get_data sg attributesK of
+ Attributes x => x
+ | _ => sys_error "get_attributes_sg");
+
+val get_attributes = get_attributes_sg o Theory.sign_of;
+
+
+(* get global / local attributes *)
+
+fun attrs which thy raw_name args =
+ let
+ val {space, attrs} = get_attributes thy;
+ val name = NameSpace.intern space raw_name;
+ in
+ (case Symtab.lookup (attrs, name) of
+ None => error ("Unknown attribute: " ^ quote name)
+ | Some p => which p args)
+ end;
+
+val global_attr = attrs fst;
+val local_attr = attrs snd;
+
+
+(* add_attrs *)
+
+fun add_attrs raw_attrs thy =
+ let
+ val full = Sign.full_name (Theory.sign_of thy);
+ val new_attrs = map (apfst full) raw_attrs;
+
+ val {space, attrs} = get_attributes thy;
+ val space' = NameSpace.extend (space, map fst new_attrs);
+ val attrs' = Symtab.extend (attrs, new_attrs)
+ handle Symtab.DUPS dups =>
+ error ("Duplicate declaration of attributes(s) " ^ commas_quote dups);
+ in
+ Theory.put_data (attributesK, Attributes {space = space', attrs = attrs'}) thy
+ end;
+
+
+end;
+
+
+structure BasicAttribute: BASIC_ATTRIBUTE = Attribute;
+open BasicAttribute;