src/Pure/Isar/attrib.ML
author wenzelm
Wed, 09 Aug 2006 00:12:38 +0200
changeset 20366 867696dc64fc
parent 20334 60157137a0eb
child 20906 63150f3a103d
permissions -rw-r--r--
global goals/qeds: after_qed operates on Proof.context (potentially local_theory); theorem/interpretation: slightly more uniform treatment of after_qeds; theorem conclusion: proper fix_frees;
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
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
     5
Symbolic representation of attributes -- with name and syntax.
5823
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
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    11
  val Attribute: bstring -> (Args.src -> attribute) -> string -> unit
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    12
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    13
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    14
signature ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    15
sig
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    16
  include BASIC_ATTRIB
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    17
  type src
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    18
  exception ATTRIB_FAIL of (string * Position.T) * exn
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    19
  val intern: theory -> xstring -> string
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    20
  val intern_src: theory -> src -> src
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    21
  val attribute: theory -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    22
  val attribute_i: theory -> src -> attribute
18905
5542716503ab more generic type for map_specs/facts;
wenzelm
parents: 18872
diff changeset
    23
  val map_specs: ('a -> 'att) ->
5542716503ab more generic type for map_specs/facts;
wenzelm
parents: 18872
diff changeset
    24
    (('c * 'a list) * 'd) list -> (('c * 'att list) * 'd) list
5542716503ab more generic type for map_specs/facts;
wenzelm
parents: 18872
diff changeset
    25
  val map_facts: ('a -> 'att) ->
17105
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
    26
    (('c * 'a list) * ('d * 'a list) list) list ->
18905
5542716503ab more generic type for map_specs/facts;
wenzelm
parents: 18872
diff changeset
    27
    (('c * 'att list) * ('d * 'att list) list) list
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20241
diff changeset
    28
  val crude_closure: Proof.context -> src -> src
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    29
  val add_attributes: (bstring * (src -> attribute) * string) list -> theory -> theory
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    30
  val thm: Context.generic * Args.T list -> thm * (Context.generic * Args.T list)
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    31
  val thms: Context.generic * Args.T list -> thm list * (Context.generic * Args.T list)
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
    32
  val multi_thm: Context.generic * Args.T list -> thm list * (Context.generic * Args.T list)
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    33
  val syntax: (Context.generic * Args.T list ->
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    34
    attribute * (Context.generic * Args.T list)) -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    35
  val no_args: attribute -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    36
  val add_del_args: attribute -> attribute -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    37
  val internal: attribute -> src
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    38
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    39
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    40
structure Attrib: ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    41
struct
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    42
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    43
type src = Args.src;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    44
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    45
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    46
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    47
(** named attributes **)
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    48
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    49
(* theory data *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    50
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    51
structure AttributesData = TheoryDataFun
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    52
(struct
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    53
  val name = "Isar/attributes";
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    54
  type T = (((src -> attribute) * string) * stamp) NameSpace.table;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    55
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    56
  val empty = NameSpace.empty_table;
6546
995a66249a9b theory data: copy;
wenzelm
parents: 6448
diff changeset
    57
  val copy = I;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    58
  val extend = I;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    59
17496
26535df536ae slight adaptions to library changes
haftmann
parents: 17412
diff changeset
    60
  fun merge _ tables = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUPS dups =>
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    61
    error ("Attempt to merge different versions of attribute(s) " ^ commas_quote dups);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    62
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    63
  fun print _ attribs =
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    64
    let
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    65
      fun prt_attr (name, ((_, comment), _)) = Pretty.block
6846
f2380295d4dd cond_extern_table;
wenzelm
parents: 6772
diff changeset
    66
        [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    67
    in
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    68
      [Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))]
9216
0842edfc8245 removed help_attributes;
wenzelm
parents: 9008
diff changeset
    69
      |> Pretty.chunks |> Pretty.writeln
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    70
    end;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    71
end);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    72
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18678
diff changeset
    73
val _ = Context.add_setup AttributesData.init;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    74
val print_attributes = AttributesData.print;
7611
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7598
diff changeset
    75
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    76
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    77
(* interning *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    78
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    79
val intern = NameSpace.intern o #1 o AttributesData.get;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    80
val intern_src = Args.map_name o intern;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    81
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    82
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    83
(* get attributes *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    84
5912
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    85
exception ATTRIB_FAIL of (string * Position.T) * exn;
3f95adea10c0 exception ATTRIB_FAIL;
wenzelm
parents: 5894
diff changeset
    86
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    87
fun attribute_i thy =
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    88
  let
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    89
    val attrs = #2 (AttributesData.get thy);
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    90
    fun attr src =
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    91
      let val ((name, _), pos) = Args.dest_src src in
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17221
diff changeset
    92
        (case Symtab.lookup attrs name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
    93
          NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    94
        | SOME ((att, _), _) => transform_failure (curry ATTRIB_FAIL (name, pos)) (att src))
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    95
      end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    96
  in attr end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    97
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    98
fun attribute thy = attribute_i thy o intern_src thy;
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    99
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   100
17105
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   101
(* attributed declarations *)
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   102
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   103
fun map_specs f = map (apfst (apsnd (map f)));
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   104
fun map_facts f = map (apfst (apsnd (map f)) o apsnd (map (apsnd (map f))));
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   105
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   106
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   107
(* crude_closure *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   108
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   109
(*Produce closure without knowing facts in advance! The following
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   110
  works reasonably well for attribute parsers that do not peek at the
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   111
  thm structure.*)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   112
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   113
fun crude_closure ctxt src =
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   114
 (try (fn () => attribute_i (ProofContext.theory_of ctxt) src
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   115
    (Context.Proof ctxt, Drule.asm_rl)) ();
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   116
  Args.closure src);
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   117
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   118
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   119
(* add_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   120
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   121
fun add_attributes raw_attrs thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   122
  let
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   123
    val new_attrs =
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   124
      raw_attrs |> map (fn (name, att, comment) => (name, ((att, comment), stamp ())));
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
   125
    fun add attrs = NameSpace.extend_table (Sign.naming_of thy) (attrs, new_attrs)
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
   126
      handle Symtab.DUPS dups =>
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
   127
        error ("Duplicate declaration of attributes(s) " ^ commas_quote dups);
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
   128
  in AttributesData.map add thy end;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   129
5879
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
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   143
local
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   144
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   145
val get_thms = Context.cases PureThy.get_thms ProofContext.get_thms;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   146
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   147
fun gen_thm pick = Scan.depend (fn st =>
20029
16957e34cfab thm parsers: include Args.internal_fact;
wenzelm
parents: 19798
diff changeset
   148
 (Args.internal_fact
16957e34cfab thm parsers: include Args.internal_fact;
wenzelm
parents: 19798
diff changeset
   149
    >> (fn fact => ("<fact>", Name "", fact)) ||
16957e34cfab thm parsers: include Args.internal_fact;
wenzelm
parents: 19798
diff changeset
   150
  Scan.ahead Args.alt_name -- Args.named_fact (get_thms st o Fact)
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   151
    >> (fn (s, fact) => ("", Fact s, fact)) ||
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   152
  Scan.ahead Args.name -- Args.named_fact (get_thms st o Name) -- Args.thm_sel
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   153
    >> (fn ((name, fact), sel) => (name, NameSelection (name, sel), fact)) ||
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   154
  Scan.ahead Args.name -- Args.named_fact (get_thms st o Name)
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   155
    >> (fn (name, fact) => (name, Name name, fact))) --
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   156
  Args.opt_attribs (intern (Context.theory_of st))
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   157
  >> (fn ((name, thmref, fact), srcs) =>
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   158
    let
16498
9d265401fee0 thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   159
      val ths = PureThy.select_thm thmref fact;
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   160
      val atts = map (attribute_i (Context.theory_of st)) srcs;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   161
      val (st', ths') = foldl_map (Library.apply atts) (st, ths);
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   162
    in (st', pick name ths') end));
15456
956d6acacf89 Specific theorems in a named list of theorems can now be referred to
berghofe
parents: 15117
diff changeset
   163
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   164
in
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   165
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   166
val thm = gen_thm PureThy.single_thm;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   167
val multi_thm = gen_thm (K I);
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19473
diff changeset
   168
val thms = Scan.repeat multi_thm >> flat;
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   169
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   170
end;
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   171
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   172
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   173
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   174
(** attribute syntax **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   175
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   176
fun syntax scan src (st, th) =
8282
58a33fd5b30c tuned syntax wrapper;
wenzelm
parents: 8164
diff changeset
   177
  let val (st', f) = Args.syntax "attribute" scan src st
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   178
  in f (st', th) end;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   179
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   180
fun no_args x = syntax (Scan.succeed x);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   181
10034
4bca6b2d2589 tuned args;
wenzelm
parents: 9952
diff changeset
   182
fun add_del_args add del x = syntax
4bca6b2d2589 tuned args;
wenzelm
parents: 9952
diff changeset
   183
  (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
   184
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   185
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   186
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   187
(** basic attributes **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   188
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   189
(* tags *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   190
18799
f137c5e971f5 moved theorem tags from Drule to PureThy;
wenzelm
parents: 18734
diff changeset
   191
fun tagged x = syntax (tag >> PureThy.tag) x;
f137c5e971f5 moved theorem tags from Drule to PureThy;
wenzelm
parents: 18734
diff changeset
   192
fun untagged x = syntax (Scan.lift Args.name >> PureThy.untag) x;
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   193
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   194
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   195
(* rule composition *)
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   196
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   197
val COMP_att =
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   198
  syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   199
    >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => Drule.compose_single (A, i, B))));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   200
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   201
val THEN_att =
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   202
  syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   203
    >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B))));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   204
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   205
val OF_att =
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   206
  syntax (thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A)));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   207
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   208
13782
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   209
(* rename_abs *)
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   210
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   211
fun rename_abs src = syntax
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   212
  (Scan.lift (Scan.repeat (Args.maybe Args.name) >> (apsnd o Drule.rename_bvars'))) src;
13782
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   213
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   214
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   215
(* unfold / fold definitions *)
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   216
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   217
fun unfolded_syntax rule =
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   218
  syntax (thms >>
18872
020044cf1510 (un)folded: removed '(raw)' option;
wenzelm
parents: 18839
diff changeset
   219
    (fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths)));
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   220
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   221
val unfolded = unfolded_syntax LocalDefs.unfold;
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   222
val folded = unfolded_syntax LocalDefs.fold;
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   223
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   224
8368
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   225
(* rule cases *)
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   226
10528
a0483268262d added "consumes" attribute;
wenzelm
parents: 10151
diff changeset
   227
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
   228
fun case_names x = syntax (Scan.lift (Scan.repeat1 Args.name) >> RuleCases.case_names) x;
18236
dd445f5cb28e added case_conclusion attribute;
wenzelm
parents: 18037
diff changeset
   229
fun case_conclusion x =
dd445f5cb28e added case_conclusion attribute;
wenzelm
parents: 18037
diff changeset
   230
  syntax (Scan.lift (Args.name -- Scan.repeat Args.name) >> RuleCases.case_conclusion) x;
8368
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   231
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
   232
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   233
20241
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   234
(* rule format *)
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   235
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   236
fun rule_format_att x = syntax (Args.mode "no_asm"
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   237
  >> (fn true => ObjectLogic.rule_format_no_asm | false => ObjectLogic.rule_format)) x;
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   238
20241
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   239
fun elim_format x = no_args (Thm.rule_attribute (K Tactic.make_elim)) x;
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   240
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   241
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   242
(* misc rules *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   243
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   244
fun standard x = no_args (Thm.rule_attribute (K Drule.standard)) x;
20241
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   245
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   246
fun no_vars x = no_args (Thm.rule_attribute (fn ctxt => fn th =>
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   247
  let val ((_, [th']), _) = Variable.import true [th] (Context.proof_of ctxt)
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   248
  in th' end)) x;
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   249
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   250
fun eta_long x = no_args (Thm.rule_attribute (K (Drule.fconv_rule Drule.eta_long_conversion))) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   251
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   252
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   253
(* internal attribute *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   254
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   255
fun internal att = Args.src (("Pure.attribute", [Args.mk_attribute att]), Position.none);
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   256
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   257
fun internal_att x = syntax (Scan.lift Args.internal_attribute) x;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   258
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   259
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   260
(* theory setup *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   261
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   262
val _ = Context.add_setup
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18678
diff changeset
   263
 (add_attributes
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   264
   [("tagged", tagged, "tagged theorem"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   265
    ("untagged", untagged, "untagged theorem"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   266
    ("COMP", COMP_att, "direct composition with rules (no lifting)"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   267
    ("THEN", THEN_att, "resolution with rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   268
    ("OF", OF_att, "rule applied to facts"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   269
    ("rename_abs", rename_abs, "rename bound variables in abstractions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   270
    ("unfolded", unfolded, "unfolded definitions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   271
    ("folded", folded, "folded definitions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   272
    ("standard", standard, "result put into standard form"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   273
    ("elim_format", elim_format, "destruct rule turned into elimination rule format"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   274
    ("no_vars", no_vars, "frozen schematic vars"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   275
    ("eta_long", eta_long, "put theorem into eta long beta normal form"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   276
    ("consumes", consumes, "number of consumed facts"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   277
    ("case_names", case_names, "named rule cases"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   278
    ("case_conclusion", case_conclusion, "named conclusion of rule cases"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   279
    ("params", params, "named rule parameters"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   280
    ("atomize", no_args ObjectLogic.declare_atomize, "declaration of atomize rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   281
    ("rulify", no_args ObjectLogic.declare_rulify, "declaration of rulify rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   282
    ("rule_format", rule_format_att, "result put into standard rule format"),
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   283
    ("defn", add_del_args LocalDefs.defn_add LocalDefs.defn_del,
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   284
      "declaration of definitional transformations"),
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   285
    ("attribute", internal_att, "internal attribute")]);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   286
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   287
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   288
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   289
structure BasicAttrib: BASIC_ATTRIB = Attrib;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   290
open BasicAttrib;