src/Pure/Isar/attrib.ML
author kleing
Mon, 21 Jun 2004 10:25:57 +0200
changeset 14981 e73f8140af78
parent 14718 f52f2cf2d137
child 15117 b860e444c1db
permissions -rw-r--r--
Merged in license change from Isabelle2004
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
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    11
  val Attribute: bstring -> (Args.src -> theory attribute) * (Args.src -> Proof.context attribute)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    12
    -> string -> unit
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    13
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    14
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    15
signature ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    16
sig
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    17
  include BASIC_ATTRIB
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    18
  exception ATTRIB_FAIL of (string * Position.T) * exn
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    19
  val global_attribute: theory -> Args.src -> theory attribute
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    20
  val local_attribute: theory -> Args.src -> Proof.context attribute
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    21
  val local_attribute': Proof.context -> Args.src -> Proof.context attribute
7673
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
    22
  val undef_global_attribute: theory attribute
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
    23
  val undef_local_attribute: Proof.context attribute
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    24
  val add_attributes: (bstring * ((Args.src -> theory attribute) *
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    25
      (Args.src -> Proof.context attribute)) * string) list -> theory -> theory
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
    26
  val global_thm: theory * Args.T list -> thm * (theory * Args.T list)
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
    27
  val global_thms: theory * Args.T list -> thm list * (theory * Args.T list)
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
    28
  val global_thmss: theory * Args.T list -> thm list * (theory * Args.T list)
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
    29
  val local_thm: Proof.context * Args.T list -> thm * (Proof.context * Args.T list)
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
    30
  val local_thms: Proof.context * Args.T list -> thm list * (Proof.context * Args.T list)
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
    31
  val local_thmss: Proof.context * Args.T list -> thm list * (Proof.context * Args.T list)
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    32
  val syntax: ('a * Args.T list -> 'a attribute * ('a * Args.T list)) -> Args.src -> 'a attribute
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    33
  val no_args: 'a attribute -> Args.src -> 'a attribute
8633
427ead639d8a added add_del_args;
wenzelm
parents: 8496
diff changeset
    34
  val add_del_args: 'a attribute -> 'a attribute -> Args.src -> 'a attribute
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    35
  val setup: (theory -> theory) list
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    36
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    37
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    38
structure Attrib: ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    39
struct
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    40
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    41
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    42
(** attributes theory data **)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    43
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    44
(* data kind 'Isar/attributes' *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    45
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    46
structure AttributesDataArgs =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    47
struct
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    48
  val name = "Isar/attributes";
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    49
  type T =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    50
    {space: NameSpace.T,
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    51
     attrs:
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    52
       ((((Args.src -> theory attribute) * (Args.src -> Proof.context attribute))
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    53
         * string) * stamp) Symtab.table};
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    54
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    55
  val empty = {space = NameSpace.empty, attrs = Symtab.empty};
6546
995a66249a9b theory data: copy;
wenzelm
parents: 6448
diff changeset
    56
  val copy = I;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    57
  val prep_ext = I;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    58
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    59
  fun merge ({space = space1, attrs = attrs1}, {space = space2, attrs = attrs2}) =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    60
    {space = NameSpace.merge (space1, space2),
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    61
      attrs = Symtab.merge eq_snd (attrs1, attrs2) handle Symtab.DUPS dups =>
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    62
        error ("Attempt to merge different versions of attributes " ^ commas_quote dups)};
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    63
9216
0842edfc8245 removed help_attributes;
wenzelm
parents: 9008
diff changeset
    64
  fun print _ {space, attrs} =
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    65
    let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    66
      fun prt_attr (name, ((_, comment), _)) = Pretty.block
6846
f2380295d4dd cond_extern_table;
wenzelm
parents: 6772
diff changeset
    67
        [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    68
    in
8720
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8687
diff changeset
    69
      [Pretty.big_list "attributes:" (map prt_attr (NameSpace.cond_extern_table space attrs))]
9216
0842edfc8245 removed help_attributes;
wenzelm
parents: 9008
diff changeset
    70
      |> Pretty.chunks |> Pretty.writeln
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    71
    end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    72
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    73
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    74
structure AttributesData = TheoryDataFun(AttributesDataArgs);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    75
val print_attributes = AttributesData.print;
7611
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7598
diff changeset
    76
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    77
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    78
(* get global / local attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    79
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    80
exception ATTRIB_FAIL of (string * Position.T) * exn;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    81
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    82
fun gen_attribute which thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    83
  let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    84
    val {space, attrs} = AttributesData.get thy;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    85
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    86
    fun attr src =
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    87
      let
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    88
        val ((raw_name, _), pos) = Args.dest_src src;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    89
        val name = NameSpace.intern space raw_name;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    90
      in
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    91
        (case Symtab.lookup (attrs, name) of
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    92
          None => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    93
        | Some ((p, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (which p src))
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    94
      end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    95
  in attr end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    96
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    97
val global_attribute = gen_attribute fst;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    98
val local_attribute = gen_attribute snd;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    99
val local_attribute' = local_attribute o ProofContext.theory_of;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   100
7673
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   101
val undef_global_attribute: theory attribute =
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   102
  fn _ => error "attribute undefined in theory context";
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   103
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   104
val undef_local_attribute: Proof.context attribute =
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   105
  fn _ => error "attribute undefined in proof context";
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   106
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   107
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   108
(* add_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   109
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   110
fun add_attributes raw_attrs thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   111
  let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   112
    val full = Sign.full_name (Theory.sign_of thy);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   113
    val new_attrs =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   114
      map (fn (name, (f, g), comment) => (full name, (((f, g), comment), stamp ()))) raw_attrs;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   115
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   116
    val {space, attrs} = AttributesData.get thy;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   117
    val space' = NameSpace.extend (space, map fst new_attrs);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   118
    val attrs' = Symtab.extend (attrs, new_attrs) handle Symtab.DUPS dups =>
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   119
      error ("Duplicate declaration of attributes(s) " ^ commas_quote dups);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   120
  in thy |> AttributesData.put {space = space', attrs = attrs'} end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   121
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   122
(*implicit version*)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   123
fun Attribute name att cmt = Context.>> (add_attributes [(name, att, cmt)]);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   124
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   125
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   126
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   127
(** attribute parsers **)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   128
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   129
(* tags *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   130
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   131
fun tag x = Scan.lift (Args.name -- Scan.repeat Args.name) x;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   132
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   133
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   134
(* theorems *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   135
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   136
fun gen_thm get attrib app =
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   137
  Scan.depend (fn st => Args.name -- Args.opt_attribs >>
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   138
    (fn (name, srcs) => app ((st, get st name), map (attrib st) srcs)));
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   139
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   140
val global_thm = gen_thm PureThy.get_thm global_attribute Thm.apply_attributes;
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   141
val global_thms = gen_thm PureThy.get_thms global_attribute Thm.applys_attributes;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   142
val global_thmss = Scan.repeat global_thms >> flat;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   143
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   144
val local_thm = gen_thm ProofContext.get_thm local_attribute' Thm.apply_attributes;
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   145
val local_thms = gen_thm ProofContext.get_thms local_attribute' Thm.applys_attributes;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   146
val local_thmss = Scan.repeat local_thms >> flat;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   147
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   148
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   149
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   150
(** attribute syntax **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   151
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   152
fun syntax scan src (st, th) =
8282
58a33fd5b30c tuned syntax wrapper;
wenzelm
parents: 8164
diff changeset
   153
  let val (st', f) = Args.syntax "attribute" scan src st
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   154
  in f (st', th) end;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   155
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   156
fun no_args x = syntax (Scan.succeed x);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   157
10034
4bca6b2d2589 tuned args;
wenzelm
parents: 9952
diff changeset
   158
fun add_del_args add del x = syntax
4bca6b2d2589 tuned args;
wenzelm
parents: 9952
diff changeset
   159
  (Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add)) x;
8633
427ead639d8a added add_del_args;
wenzelm
parents: 8496
diff changeset
   160
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   161
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   162
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   163
(** Pure attributes **)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   164
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   165
(* tags *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   166
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   167
fun gen_tagged x = syntax (tag >> Drule.tag) x;
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   168
fun gen_untagged x = syntax (Scan.lift Args.name >> Drule.untag) x;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   169
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   170
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   171
(* COMP *)
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   172
6948
01f3c7866ead COMP: optional position;
wenzelm
parents: 6933
diff changeset
   173
fun comp (i, B) (x, A) = (x, Drule.compose_single (A, i, B));
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   174
10151
631628d6dd03 'THEN', 'COMP': improved optional position arg;
wenzelm
parents: 10034
diff changeset
   175
fun gen_COMP thm = syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm >> comp);
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   176
val COMP_global = gen_COMP global_thm;
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   177
val COMP_local = gen_COMP local_thm;
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   178
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   179
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   180
(* RS *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   181
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   182
fun resolve (i, B) (x, A) = (x, A RSN (i, B));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   183
10151
631628d6dd03 'THEN', 'COMP': improved optional position arg;
wenzelm
parents: 10034
diff changeset
   184
fun gen_RS thm = syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm >> resolve);
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   185
val RS_global = gen_RS global_thm;
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   186
val RS_local = gen_RS local_thm;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   187
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   188
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   189
(* OF *)
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   190
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   191
fun apply Bs (x, A) = (x, Bs MRS A);
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   192
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   193
val OF_global = syntax (global_thmss >> apply);
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   194
val OF_local = syntax (local_thmss >> apply);
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   195
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   196
14718
wenzelm
parents: 14287
diff changeset
   197
(* where *)
wenzelm
parents: 14287
diff changeset
   198
wenzelm
parents: 14287
diff changeset
   199
(*named instantiations; permits instantiation of type and term variables*)
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   200
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   201
fun read_instantiate _ [] _ thm = thm
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   202
  | read_instantiate context_of insts x thm =
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   203
      let
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   204
        val ctxt = context_of x;
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   205
        val sign = ProofContext.sign_of ctxt;
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   206
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   207
        (* Separate type and term insts,
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   208
           type insts must occur strictly before term insts *)
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   209
        fun has_type_var ((x, _), _) = (case Symbol.explode x of
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   210
             "'"::cs => true | cs => false);
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   211
        val (Tinst, tinsts) = take_prefix has_type_var insts;
14718
wenzelm
parents: 14287
diff changeset
   212
        val _ = if exists has_type_var tinsts
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   213
              then error
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   214
                "Type instantiations must occur before term instantiations."
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   215
              else ();
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   216
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   217
        val Tinsts = filter has_type_var insts;
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   218
        val tinsts = filter_out has_type_var insts;
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   219
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   220
        (* Type instantiations first *)
14718
wenzelm
parents: 14287
diff changeset
   221
        (* Process type insts: Tinsts_env *)
wenzelm
parents: 14287
diff changeset
   222
        fun absent xi = error
wenzelm
parents: 14287
diff changeset
   223
              ("No such type variable in theorem: " ^
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   224
               Syntax.string_of_vname xi);
14718
wenzelm
parents: 14287
diff changeset
   225
        val (rtypes, rsorts) = types_sorts thm;
wenzelm
parents: 14287
diff changeset
   226
        fun readT (xi, s) =
wenzelm
parents: 14287
diff changeset
   227
            let val S = case rsorts xi of Some S => S | None => absent xi;
wenzelm
parents: 14287
diff changeset
   228
                val T = ProofContext.read_typ ctxt s;
wenzelm
parents: 14287
diff changeset
   229
            in if Sign.typ_instance sign (T, TVar (xi, S)) then (xi, T)
wenzelm
parents: 14287
diff changeset
   230
               else error
wenzelm
parents: 14287
diff changeset
   231
                 ("Instantiation of " ^ Syntax.string_of_vname xi ^ " fails")
wenzelm
parents: 14287
diff changeset
   232
            end;
wenzelm
parents: 14287
diff changeset
   233
        val Tinsts_env = map readT Tinsts;
wenzelm
parents: 14287
diff changeset
   234
        val cenvT = map (apsnd (Thm.ctyp_of sign)) (Tinsts_env);
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   235
        val thm' = Thm.instantiate (cenvT, []) thm;
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   236
           (* Instantiate, but don't normalise: *)
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   237
           (* this happens after term insts anyway. *)
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   238
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   239
        (* Term instantiations *)
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   240
        val vars = Drule.vars_of thm';
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   241
        fun get_typ xi =
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   242
          (case assoc (vars, xi) of
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   243
            Some T => T
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   244
          | None => error ("No such variable in theorem: " ^ Syntax.string_of_vname xi));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   245
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   246
        val (xs, ss) = Library.split_list tinsts;
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   247
        val Ts = map get_typ xs;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   248
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   249
        val used = add_term_tvarnames (prop_of thm',[])
14718
wenzelm
parents: 14287
diff changeset
   250
        (* Names of TVars occuring in thm'.  read_def_termTs ensures
14257
a7ef3f7588c5 Type inference bug in Isar attributes "where" and "of" fixed.
ballarin
parents: 14082
diff changeset
   251
           that new TVars introduced when reading the instantiation string
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   252
           are distinct from those occuring in the theorem. *)
14257
a7ef3f7588c5 Type inference bug in Isar attributes "where" and "of" fixed.
ballarin
parents: 14082
diff changeset
   253
14718
wenzelm
parents: 14287
diff changeset
   254
        val (ts, envT) =
wenzelm
parents: 14287
diff changeset
   255
          ProofContext.read_termTs ctxt (K false) (K None) (K None) used (ss ~~ Ts);
wenzelm
parents: 14287
diff changeset
   256
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   257
        val cenvT = map (apsnd (Thm.ctyp_of sign)) envT;
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   258
        val cenv =
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   259
          map (fn (xi, t) => pairself (Thm.cterm_of sign) (Var (xi, fastype_of t), t))
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   260
            (gen_distinct (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) (xs ~~ ts));
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   261
      in
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   262
        thm'
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   263
        |> Drule.instantiate (cenvT, cenv)
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   264
        |> RuleCases.save thm
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   265
      end;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   266
6448
932f27366c8f and_list;
wenzelm
parents: 6091
diff changeset
   267
fun insts x = Args.and_list (Scan.lift (Args.var --| Args.$$$ "=" -- Args.name)) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   268
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   269
fun gen_where context_of = syntax (insts >> (Drule.rule_attribute o read_instantiate context_of));
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   270
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   271
val where_global = gen_where ProofContext.init;
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   272
val where_local = gen_where I;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   273
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   274
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   275
(* of: positional instantiations *)
14287
f630017ed01c Isar: where attribute supports instantiation of type variables.
ballarin
parents: 14257
diff changeset
   276
(*        permits instantiation of term variables only *)
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   277
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   278
fun read_instantiate' _ ([], []) _ thm = thm
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   279
  | read_instantiate' context_of (args, concl_args) x thm =
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   280
      let
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   281
        fun zip_vars _ [] = []
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   282
          | zip_vars (_ :: xs) (None :: opt_ts) = zip_vars xs opt_ts
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   283
          | zip_vars ((x, _) :: xs) (Some t :: opt_ts) = (x, t) :: zip_vars xs opt_ts
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   284
          | zip_vars [] _ = error "More instantiations than variables in theorem";
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   285
        val insts =
12804
163a85ba885b Tactic.norm_hhf renamed to Tactic.norm_hhf_rule;
wenzelm
parents: 12775
diff changeset
   286
          zip_vars (Drule.vars_of_terms [Thm.prop_of thm]) args @
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   287
          zip_vars (Drule.vars_of_terms [Thm.concl_of thm]) concl_args;
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   288
      in
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   289
        thm
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   290
        |> read_instantiate context_of insts x
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   291
        |> RuleCases.save thm
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   292
      end;
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   293
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   294
val concl = Args.$$$ "concl" -- Args.colon;
8687
24bff69370f0 Args.name_dummy;
wenzelm
parents: 8633
diff changeset
   295
val inst_arg = Scan.unless concl Args.name_dummy;
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   296
val inst_args = Scan.repeat inst_arg;
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   297
fun insts' x = (inst_args -- Scan.optional (concl |-- Args.!!! inst_args) []) x;
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   298
10807
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   299
fun gen_of context_of =
ae001d5119fc export read_inst', inst';
wenzelm
parents: 10528
diff changeset
   300
  syntax (Scan.lift insts' >> (Drule.rule_attribute o read_instantiate' context_of));
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   301
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   302
val of_global = gen_of ProofContext.init;
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   303
val of_local = gen_of I;
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   304
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   305
13782
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   306
(* rename_abs *)
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   307
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   308
fun rename_abs src = syntax
14082
c69d5bf3047d Moved function for renaming bound variables to Pure/drule.ML
berghofe
parents: 13782
diff changeset
   309
  (Scan.lift (Scan.repeat Args.name_dummy >> (apsnd o Drule.rename_bvars'))) src;
13782
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   310
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   311
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   312
(* unfold / fold definitions *)
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   313
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   314
fun gen_rewrite rew defs (x, thm) = (x, rew defs thm);
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   315
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   316
val unfolded_global = syntax (global_thmss >> gen_rewrite Tactic.rewrite_rule);
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   317
val unfolded_local = syntax (local_thmss >> gen_rewrite Tactic.rewrite_rule);
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   318
val folded_global = syntax (global_thmss >> gen_rewrite Tactic.fold_rule);
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   319
val folded_local = syntax (local_thmss >> gen_rewrite Tactic.fold_rule);
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   320
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   321
8368
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   322
(* rule cases *)
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   323
10528
a0483268262d added "consumes" attribute;
wenzelm
parents: 10151
diff changeset
   324
fun consumes x = syntax (Scan.lift (Scan.optional Args.nat 1) >> RuleCases.consumes) x;
8368
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   325
fun case_names x = syntax (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names) x;
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   326
fun params x = syntax (Args.and_list1 (Scan.lift (Scan.repeat Args.name)) >> RuleCases.params) x;
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   327
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   328
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   329
(* rule_format *)
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   330
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   331
fun rule_format_att x = syntax
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   332
  (Scan.lift (Args.parens (Args.$$$ "no_asm")
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   333
  >> K ObjectLogic.rule_format_no_asm || Scan.succeed ObjectLogic.rule_format)) x;
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   334
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   335
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   336
(* misc rules *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   337
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   338
fun standard x = no_args (Drule.rule_attribute (K Drule.standard)) x;
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9902
diff changeset
   339
fun elim_format x = no_args (Drule.rule_attribute (K Tactic.make_elim)) x;
9216
0842edfc8245 removed help_attributes;
wenzelm
parents: 9008
diff changeset
   340
fun no_vars x = no_args (Drule.rule_attribute (K (#1 o Drule.freeze_thaw))) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   341
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   342
13370
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   343
(* rule declarations *)
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   344
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   345
local
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   346
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   347
fun add_args a b c x = syntax
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   348
  (Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) -- (Scan.option Args.nat))
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   349
    >> (fn (f, n) => f n)) x;
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   350
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   351
fun del_args att = syntax (Scan.lift Args.del >> K att);
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   352
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   353
open ContextRules;
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   354
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   355
in
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   356
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   357
val rule_atts =
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   358
 [("intro",
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   359
   (add_args intro_bang_global intro_global intro_query_global,
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   360
    add_args intro_bang_local intro_local intro_query_local),
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   361
    "declaration of introduction rule"),
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   362
  ("elim",
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   363
   (add_args elim_bang_global elim_global elim_query_global,
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   364
    add_args elim_bang_local elim_local elim_query_local),
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   365
    "declaration of elimination rule"),
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   366
  ("dest",
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   367
   (add_args dest_bang_global dest_global dest_query_global,
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   368
    add_args dest_bang_local dest_local dest_query_local),
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   369
    "declaration of destruction rule"),
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   370
  ("rule", (del_args rule_del_global, del_args rule_del_local),
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   371
    "remove declaration of intro/elim/dest rule")];
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   372
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   373
end;
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   374
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   375
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   376
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   377
(** theory setup **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   378
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   379
(* pure_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   380
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   381
val pure_attributes =
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   382
 [("tagged", (gen_tagged, gen_tagged), "tagged theorem"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   383
  ("untagged", (gen_untagged, gen_untagged), "untagged theorem"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   384
  ("COMP", (COMP_global, COMP_local), "direct composition with rules (no lifting)"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   385
  ("THEN", (RS_global, RS_local), "resolution with rule"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   386
  ("OF", (OF_global, OF_local), "rule applied to facts"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   387
  ("where", (where_global, where_local), "named instantiation of theorem"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   388
  ("of", (of_global, of_local), "rule applied to terms"),
13782
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   389
  ("rename_abs", (rename_abs, rename_abs), "rename bound variables in abstractions"),
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   390
  ("unfolded", (unfolded_global, unfolded_local), "unfolded definitions"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   391
  ("folded", (folded_global, folded_local), "folded definitions"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   392
  ("standard", (standard, standard), "result put into standard form"),
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9902
diff changeset
   393
  ("elim_format", (elim_format, elim_format), "destruct rule turned into elimination rule format"),
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   394
  ("no_vars", (no_vars, no_vars), "frozen schematic vars"),
10528
a0483268262d added "consumes" attribute;
wenzelm
parents: 10151
diff changeset
   395
  ("consumes", (consumes, consumes), "number of consumed facts"),
9902
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   396
  ("case_names", (case_names, case_names), "named rule cases"),
1ea354905d88 improved att names / msgs;
wenzelm
parents: 9630
diff changeset
   397
  ("params", (params, params), "named rule parameters"),
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   398
  ("atomize", (no_args ObjectLogic.declare_atomize, no_args undef_local_attribute),
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   399
    "declaration of atomize rule"),
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   400
  ("rulify", (no_args ObjectLogic.declare_rulify, no_args undef_local_attribute),
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   401
    "declaration of rulify rule"),
13370
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   402
  ("rule_format", (rule_format_att, rule_format_att), "result put into standard rule format")] @
3ec0d8c8beba context rules;
wenzelm
parents: 12804
diff changeset
   403
  rule_atts;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   404
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   405
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   406
(* setup *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   407
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   408
val setup = [AttributesData.init, add_attributes pure_attributes];
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   409
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   410
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   411
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   412
structure BasicAttrib: BASIC_ATTRIB = Attrib;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   413
open BasicAttrib;