src/Pure/Isar/attrib.ML
author wenzelm
Mon, 09 Nov 1998 15:32:20 +0100
changeset 5823 ee7c198a2154
child 5879 18b8f048d93a
permissions -rw-r--r--
Symbolic theorem attributes.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/attrib.ML
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     4
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     5
Symbolic theorem attributes.
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     6
*)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     7
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     8
signature BASIC_ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
     9
sig
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    10
  val print_attributes: theory -> unit
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    11
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    12
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    13
signature ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    14
sig
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    15
  include BASIC_ATTRIB
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    16
  val global_attribute: theory -> Args.src -> theory attribute
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    17
  val local_attribute: theory -> Args.src -> Proof.context attribute
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    18
  val add_attributes: (bstring * ((Args.src -> theory attribute) *
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    19
      (Args.src -> Proof.context attribute)) * string) list -> theory -> theory
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    20
  val syntax: (Args.T list -> 'a * Args.T list) -> ('a -> 'b) -> Args.src -> 'b
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    21
  val no_args: 'a attribute -> Args.src -> 'a attribute
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    22
  val thm_args: ('a -> string list -> tthm list)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    23
    -> (tthm list -> 'a attribute) -> Args.src -> 'a attribute
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    24
  val global_thm_args: (tthm list -> theory attribute) -> Args.src -> theory attribute
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    25
  val local_thm_args: (tthm list -> Proof.context attribute)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    26
    -> Args.src -> Proof.context attribute
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    27
  val setup: (theory -> theory) list
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    28
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    29
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    30
structure Attrib: ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    31
struct
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    32
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    33
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    34
(** attributes theory data **)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    35
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    36
(* data kind 'Isar/attributes' *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    37
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    38
structure AttributesDataArgs =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    39
struct
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    40
  val name = "Isar/attributes";
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    41
  type T =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    42
    {space: NameSpace.T,
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    43
     attrs:
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    44
       ((((Args.src -> theory attribute) * (Args.src -> Proof.context attribute))
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    45
         * string) * stamp) Symtab.table};
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    46
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    47
  val empty = {space = NameSpace.empty, attrs = Symtab.empty};
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    48
  val prep_ext = I;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    49
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    50
  fun merge ({space = space1, attrs = attrs1}, {space = space2, attrs = attrs2}) =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    51
    {space = NameSpace.merge (space1, space2),
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    52
      attrs = Symtab.merge eq_snd (attrs1, attrs2) handle Symtab.DUPS dups =>
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    53
        error ("Attempt to merge different versions of attributes " ^ commas_quote dups)};
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    54
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    55
  fun print _ ({space, attrs}) =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    56
    let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    57
      fun prt_attr (name, ((_, comment), _)) = Pretty.block
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    58
        [Pretty.str (NameSpace.cond_extern space name ^ ":"), Pretty.brk 2, Pretty.str comment];
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    59
    in
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    60
      Pretty.writeln (Display.pretty_name_space ("attribute name space", space));
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    61
      Pretty.writeln (Pretty.big_list "attributes:" (map prt_attr (Symtab.dest attrs)))
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    62
    end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    63
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    64
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    65
structure AttributesData = TheoryDataFun(AttributesDataArgs);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    66
val print_attributes = AttributesData.print;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    67
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    68
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    69
(* get global / local attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    70
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    71
fun gen_attribute which thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    72
  let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    73
    val {space, attrs} = AttributesData.get thy;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    74
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    75
    fun attr ((raw_name, args), pos) =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    76
      let val name = NameSpace.intern space raw_name in
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    77
        (case Symtab.lookup (attrs, name) of
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    78
          None => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    79
        | Some ((p, _), _) => which p ((name, args), pos))
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    80
      end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    81
  in attr end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    82
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    83
val global_attribute = gen_attribute fst;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    84
val local_attribute = gen_attribute snd;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    85
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    86
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    87
(* add_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    88
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    89
fun add_attributes raw_attrs thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    90
  let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    91
    val full = Sign.full_name (Theory.sign_of thy);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    92
    val new_attrs =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    93
      map (fn (name, (f, g), comment) => (full name, (((f, g), comment), stamp ()))) raw_attrs;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    94
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    95
    val {space, attrs} = AttributesData.get thy;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    96
    val space' = NameSpace.extend (space, map fst new_attrs);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    97
    val attrs' = Symtab.extend (attrs, new_attrs) handle Symtab.DUPS dups =>
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    98
      error ("Duplicate declaration of attributes(s) " ^ commas_quote dups);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    99
  in thy |> AttributesData.put {space = space', attrs = attrs'} end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   100
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   101
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   102
(* attribute syntax *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   103
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   104
val attributeN = "attribute";
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   105
fun syntax scan = Args.syntax attributeN scan;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   106
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   107
fun no_args x = syntax (Scan.succeed x) I;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   108
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   109
fun thm_args get f = syntax (Scan.repeat Args.name)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   110
  (fn names => fn (x, ths) => f (get x names) (x, ths));
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   111
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   112
fun global_thm_args f = thm_args PureThy.get_tthmss f;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   113
fun local_thm_args f = thm_args ProofContext.get_tthmss f;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   114
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   115
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   116
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   117
(** Pure attributes **)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   118
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   119
(* tags *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   120
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   121
fun gen_tag x = syntax (Args.name -- Scan.repeat Args.name) Attribute.tag x;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   122
fun gen_untag x = syntax (Args.name -- Scan.repeat Args.name) Attribute.untag x;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   123
fun gen_lemma x = no_args Attribute.tag_lemma x;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   124
fun gen_assumption x = no_args Attribute.tag_assumption x;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   125
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   126
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   127
(* resolution *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   128
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   129
fun gen_RS get = syntax Args.name
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   130
  (fn name => fn (x, (thm, tags)) => (x, (thm RS (Attribute.thm_of (get x name)), tags)));
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   131
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   132
fun RS_global x = gen_RS PureThy.get_tthm x;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   133
fun RS_local x = gen_RS ProofContext.get_tthm x;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   134
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   135
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   136
(* pure_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   137
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   138
val pure_attributes =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   139
 [("tag", (gen_tag, gen_tag), "tag theorem"),
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   140
  ("untag", (gen_untag, gen_untag), "untag theorem"),
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   141
  ("lemma", (gen_lemma, gen_lemma), "tag as lemma"),
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   142
  ("assumption", (gen_assumption, gen_assumption), "tag as assumption"),
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   143
  ("RS", (RS_global, RS_local), "resolve with rule")];
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   144
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   145
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   146
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   147
(** theory setup **)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   148
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   149
val setup = [AttributesData.init, add_attributes pure_attributes];
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   150
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   151
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   152
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   153
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   154
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   155
structure BasicAttrib: BASIC_ATTRIB = Attrib;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   156
open BasicAttrib;