src/Pure/Isar/attrib.ML
author wenzelm
Wed, 08 Mar 2000 17:52:38 +0100
changeset 8368 bdc3ee0d8cb6
parent 8282 58a33fd5b30c
child 8496 7e4a466b18d5
permissions -rw-r--r--
added 'case_names' and 'params';
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
7611
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7598
diff changeset
    18
  val help_attributes: theory option -> unit
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    19
  exception ATTRIB_FAIL of (string * Position.T) * exn
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    20
  val global_attribute: theory -> Args.src -> theory attribute
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    21
  val local_attribute: theory -> Args.src -> Proof.context attribute
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    22
  val local_attribute': Proof.context -> Args.src -> Proof.context attribute
7673
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
    23
  val undef_global_attribute: theory attribute
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
    24
  val undef_local_attribute: Proof.context attribute
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    25
  val add_attributes: (bstring * ((Args.src -> theory attribute) *
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    26
      (Args.src -> Proof.context attribute)) * string) list -> theory -> theory
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
    27
  val global_thm: theory * Args.T list -> thm * (theory * Args.T list)
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
    28
  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
    29
  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
    30
  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
    31
  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
    32
  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
    33
  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
    34
  val no_args: 'a attribute -> Args.src -> 'a attribute
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
7367
a79d4683fadf print_help;
wenzelm
parents: 6948
diff changeset
    64
  fun print_atts verbose ({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
7367
a79d4683fadf print_help;
wenzelm
parents: 6948
diff changeset
    69
      if not verbose then ()
a79d4683fadf print_help;
wenzelm
parents: 6948
diff changeset
    70
      else Pretty.writeln (Display.pretty_name_space ("attribute name space", space));
6846
f2380295d4dd cond_extern_table;
wenzelm
parents: 6772
diff changeset
    71
      Pretty.writeln (Pretty.big_list "attributes:"
f2380295d4dd cond_extern_table;
wenzelm
parents: 6772
diff changeset
    72
        (map prt_attr (NameSpace.cond_extern_table space attrs)))
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    73
    end;
7367
a79d4683fadf print_help;
wenzelm
parents: 6948
diff changeset
    74
a79d4683fadf print_help;
wenzelm
parents: 6948
diff changeset
    75
   fun print _ = print_atts true;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    76
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    77
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    78
structure AttributesData = TheoryDataFun(AttributesDataArgs);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    79
val print_attributes = AttributesData.print;
7611
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7598
diff changeset
    80
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7598
diff changeset
    81
fun help_attributes None = writeln "attributes: (unkown theory context)"
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7598
diff changeset
    82
  | help_attributes (Some thy) = AttributesDataArgs.print_atts false (AttributesData.get thy);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    83
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    84
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    85
(* get global / local attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    86
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    87
exception ATTRIB_FAIL of (string * Position.T) * exn;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    88
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    89
fun gen_attribute which thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    90
  let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    91
    val {space, attrs} = AttributesData.get thy;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    92
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    93
    fun attr src =
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    94
      let
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    95
        val ((raw_name, _), pos) = Args.dest_src src;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    96
        val name = NameSpace.intern space raw_name;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    97
      in
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    98
        (case Symtab.lookup (attrs, name) of
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    99
          None => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   100
        | Some ((p, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (which p src))
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   101
      end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   102
  in attr end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   103
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   104
val global_attribute = gen_attribute fst;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   105
val local_attribute = gen_attribute snd;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   106
val local_attribute' = local_attribute o ProofContext.theory_of;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   107
7673
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   108
val undef_global_attribute: theory attribute =
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   109
  fn _ => error "attribute undefined in theory context";
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   110
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   111
val undef_local_attribute: Proof.context attribute =
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   112
  fn _ => error "attribute undefined in proof context";
b8e7fa177d62 added undef_global_attribute, undef_local_attribute;
wenzelm
parents: 7668
diff changeset
   113
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   114
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   115
(* add_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   116
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   117
fun add_attributes raw_attrs thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   118
  let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   119
    val full = Sign.full_name (Theory.sign_of thy);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   120
    val new_attrs =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   121
      map (fn (name, (f, g), comment) => (full name, (((f, g), comment), stamp ()))) raw_attrs;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   122
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   123
    val {space, attrs} = AttributesData.get thy;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   124
    val space' = NameSpace.extend (space, map fst new_attrs);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   125
    val attrs' = Symtab.extend (attrs, new_attrs) handle Symtab.DUPS dups =>
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   126
      error ("Duplicate declaration of attributes(s) " ^ commas_quote dups);
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   127
  in thy |> AttributesData.put {space = space', attrs = attrs'} end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   128
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   129
(*implicit version*)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   130
fun Attribute name att cmt = Context.>> (add_attributes [(name, att, cmt)]);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   131
5879
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
(** attribute parsers **)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   135
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   136
(* tags *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   137
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   138
fun tag x = Scan.lift (Args.name -- Scan.repeat Args.name) x;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   139
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   140
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   141
(* theorems *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   142
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   143
fun gen_thm get attrib app =
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   144
  Scan.depend (fn st => Args.name -- Args.opt_attribs >>
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   145
    (fn (name, srcs) => app ((st, get st name), map (attrib st) srcs)));
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   146
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   147
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
   148
val global_thms = gen_thm PureThy.get_thms global_attribute Thm.applys_attributes;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   149
val global_thmss = Scan.repeat global_thms >> flat;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   150
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   151
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
   152
val local_thms = gen_thm ProofContext.get_thms local_attribute' Thm.applys_attributes;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   153
val local_thmss = Scan.repeat local_thms >> flat;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   154
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   155
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   156
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   157
(** attribute syntax **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   158
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   159
fun syntax scan src (st, th) =
8282
58a33fd5b30c tuned syntax wrapper;
wenzelm
parents: 8164
diff changeset
   160
  let val (st', f) = Args.syntax "attribute" scan src st
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   161
  in f (st', th) end;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   162
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   163
fun no_args x = syntax (Scan.succeed x);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   164
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   165
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   166
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   167
(** Pure attributes **)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   168
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   169
(* tags *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   170
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   171
fun gen_tag x = syntax (tag >> Drule.tag) x;
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   172
fun gen_untag x = syntax (tag >> Drule.untag) x;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   173
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   174
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   175
(* transfer *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   176
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   177
fun gen_transfer theory_of = no_args (Drule.rule_attribute (fn st => Thm.transfer (theory_of st)));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   178
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   179
val global_transfer = gen_transfer I;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   180
val local_transfer = gen_transfer ProofContext.theory_of;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   181
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   182
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   183
(* COMP *)
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   184
6948
01f3c7866ead COMP: optional position;
wenzelm
parents: 6933
diff changeset
   185
fun comp (i, B) (x, A) = (x, Drule.compose_single (A, i, B));
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   186
6948
01f3c7866ead COMP: optional position;
wenzelm
parents: 6933
diff changeset
   187
fun gen_COMP thm = syntax (Scan.lift (Scan.optional Args.nat 1) -- thm >> comp);
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   188
val global_COMP = gen_COMP global_thm;
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   189
val local_COMP = gen_COMP local_thm;
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   190
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   191
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   192
(* RS *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   193
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   194
fun resolve (i, B) (x, A) = (x, A RSN (i, B));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   195
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   196
fun gen_RS thm = syntax (Scan.lift (Scan.optional Args.nat 1) -- thm >> resolve);
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   197
val global_RS = gen_RS global_thm;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   198
val local_RS = gen_RS local_thm;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   199
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   200
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   201
(* APP *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   202
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   203
fun apply Bs (x, A) = (x, Bs MRS A);
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   204
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   205
val global_APP = syntax (global_thmss >> apply);
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   206
val local_APP = syntax (local_thmss >> apply);
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   207
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   208
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   209
(* where: named instantiations *)
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   210
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   211
fun read_instantiate context_of insts x thm =
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   212
  let
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   213
    val ctxt = context_of x;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   214
    val sign = ProofContext.sign_of ctxt;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   215
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   216
    val vars = Drule.vars_of thm;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   217
    fun get_typ xi =
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   218
      (case assoc (vars, xi) of
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   219
        Some T => T
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   220
      | None => error ("No such variable in theorem: " ^ Syntax.string_of_vname xi));
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   221
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   222
    val (xs, ss) = Library.split_list insts;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   223
    val Ts = map get_typ xs;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   224
7668
80c310e76c46 removed ProofContext.declare_thm;
wenzelm
parents: 7611
diff changeset
   225
    val (ts, envT) = ProofContext.read_termTs ctxt (ss ~~ Ts);
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   226
    val cenvT = map (apsnd (Thm.ctyp_of sign)) envT;
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   227
    val cenv =
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   228
      map (fn (xi, t) => pairself (Thm.cterm_of sign) (Var (xi, fastype_of t), t))
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   229
        (gen_distinct (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) (xs ~~ ts));
8164
27f14e47e2d5 Drule.instantiate;
wenzelm
parents: 7673
diff changeset
   230
  in Drule.instantiate (cenvT, cenv) thm end;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   231
6448
932f27366c8f and_list;
wenzelm
parents: 6091
diff changeset
   232
fun insts x = Args.and_list (Scan.lift (Args.var --| Args.$$$ "=" -- Args.name)) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   233
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   234
fun gen_where context_of = syntax (insts >> (Drule.rule_attribute o read_instantiate context_of));
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   235
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   236
val global_where = gen_where ProofContext.init;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   237
val local_where = gen_where I;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   238
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   239
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   240
(* with: positional instantiations *)
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   241
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   242
fun read_instantiate' context_of (args, concl_args) x thm =
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   243
  let
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   244
    fun zip_vars _ [] = []
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   245
      | zip_vars (_ :: xs) (None :: opt_ts) = zip_vars xs opt_ts
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   246
      | zip_vars ((x, _) :: xs) (Some t :: opt_ts) = (x, t) :: zip_vars xs opt_ts
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   247
      | zip_vars [] _ = error "More instantiations than variables in theorem";
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   248
    val insts =
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   249
      zip_vars (Drule.vars_of_terms [#prop (Thm.rep_thm thm)]) args @
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   250
      zip_vars (Drule.vars_of_terms [Thm.concl_of thm]) concl_args;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   251
  in read_instantiate context_of insts x thm end;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   252
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   253
val concl = Args.$$$ "concl" -- Args.$$$ ":";
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   254
val inst_arg = Scan.unless concl (Args.$$$ "_" >> K None || Args.name >> Some);
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   255
val inst_args = Scan.repeat inst_arg;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   256
fun insts' x = Scan.lift (inst_args -- Scan.optional (concl |-- Args.!!! inst_args) []) x;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   257
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   258
fun gen_with context_of = syntax (insts' >> (Drule.rule_attribute o read_instantiate' context_of));
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   259
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   260
val global_with = gen_with ProofContext.init;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   261
val local_with = gen_with I;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   262
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   263
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   264
(* unfold / fold definitions *)
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   265
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   266
fun gen_rewrite rew defs (x, thm) = (x, rew defs thm);
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   267
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   268
val global_unfold = syntax (global_thmss >> gen_rewrite Tactic.rewrite_rule);
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   269
val local_unfold = syntax (local_thmss >> gen_rewrite Tactic.rewrite_rule);
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   270
val global_fold = syntax (global_thmss >> gen_rewrite Tactic.fold_rule);
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   271
val local_fold = syntax (local_thmss >> gen_rewrite Tactic.fold_rule);
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   272
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   273
8368
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   274
(* rule cases *)
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   275
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   276
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
   277
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
   278
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   279
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   280
(* misc rules *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   281
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   282
fun standard x = no_args (Drule.rule_attribute (K Drule.standard)) x;
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5912
diff changeset
   283
fun elimify x = no_args (Drule.rule_attribute (K Tactic.make_elim)) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   284
6933
0890fde41522 'export';
wenzelm
parents: 6874
diff changeset
   285
fun global_export x = no_args (Drule.rule_attribute (Proof.export_thm o ProofContext.init)) x;
0890fde41522 'export';
wenzelm
parents: 6874
diff changeset
   286
fun local_export x = no_args (Drule.rule_attribute Proof.export_thm) x;
0890fde41522 'export';
wenzelm
parents: 6874
diff changeset
   287
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   288
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   289
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   290
(** theory setup **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   291
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   292
(* pure_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   293
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   294
val pure_attributes =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   295
 [("tag", (gen_tag, gen_tag), "tag theorem"),
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   296
  ("untag", (gen_untag, gen_untag), "untag theorem"),
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   297
  ("COMP", (global_COMP, local_COMP), "compose rules (no lifting)"),
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   298
  ("RS", (global_RS, local_RS), "resolve with rule"),
6874
747f656e04ec renamed with/APP to of/OF;
wenzelm
parents: 6846
diff changeset
   299
  ("OF", (global_APP, local_APP), "resolve with rule -- apply rule to rules"),
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
   300
  ("where", (global_where, local_where), "named instantiation of theorem"),
6874
747f656e04ec renamed with/APP to of/OF;
wenzelm
parents: 6846
diff changeset
   301
  ("of", (global_with, local_with), "positional instantiation of theorem -- apply rule to terms"),
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   302
  ("unfold", (global_unfold, local_unfold), "unfold definitions"),
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   303
  ("fold", (global_fold, local_fold), "fold definitions"),
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   304
  ("standard", (standard, standard), "put theorem into standard form"),
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   305
  ("elimify", (elimify, elimify), "turn destruct rule into elimination rule"),
8368
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   306
  ("case_names", (case_names, case_names), "name rule cases"),
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   307
  ("params", (params, params), "name rule parameters"),
6933
0890fde41522 'export';
wenzelm
parents: 6874
diff changeset
   308
  ("transfer", (global_transfer, local_transfer), "transfer theorem to this theory"),
0890fde41522 'export';
wenzelm
parents: 6874
diff changeset
   309
  ("export", (global_export, local_export), "export theorem from context")];
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   310
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   311
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   312
(* setup *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   313
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   314
val setup = [AttributesData.init, add_attributes pure_attributes];
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   315
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   316
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   317
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   318
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   319
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   320
structure BasicAttrib: BASIC_ATTRIB = Attrib;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   321
open BasicAttrib;