src/Pure/Isar/attrib.ML
author wenzelm
Sat, 28 Jul 2007 20:40:27 +0200
changeset 24022 ab76c73b3b58
parent 24008 63ff2445ce1e
child 24030 d39d64d96e71
permissions -rw-r--r--
tuned;
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
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    11
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    12
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    13
signature ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    14
sig
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    15
  include BASIC_ATTRIB
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    16
  type src
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    17
  val intern: theory -> xstring -> string
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    18
  val intern_src: theory -> src -> src
21031
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    19
  val pretty_attribs: Proof.context -> src list -> Pretty.T list
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    20
  val attribute: theory -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    21
  val attribute_i: theory -> src -> attribute
18905
5542716503ab more generic type for map_specs/facts;
wenzelm
parents: 18872
diff changeset
    22
  val map_specs: ('a -> 'att) ->
5542716503ab more generic type for map_specs/facts;
wenzelm
parents: 18872
diff changeset
    23
    (('c * 'a list) * 'd) list -> (('c * 'att list) * 'd) list
5542716503ab more generic type for map_specs/facts;
wenzelm
parents: 18872
diff changeset
    24
  val map_facts: ('a -> 'att) ->
17105
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
    25
    (('c * 'a list) * ('d * 'a list) list) list ->
18905
5542716503ab more generic type for map_specs/facts;
wenzelm
parents: 18872
diff changeset
    26
    (('c * 'att list) * ('d * 'att list) list) list
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 20241
diff changeset
    27
  val crude_closure: Proof.context -> src -> src
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    28
  val add_attributes: (bstring * (src -> attribute) * string) list -> theory -> theory
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    29
  val thm: Context.generic * Args.T list -> thm * (Context.generic * Args.T list)
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    30
  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
    31
  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
    32
  val syntax: (Context.generic * Args.T list ->
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    33
    attribute * (Context.generic * Args.T list)) -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    34
  val no_args: attribute -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    35
  val add_del_args: attribute -> attribute -> src -> attribute
21658
5e31241e1e3c Attrib.internal: morphism;
wenzelm
parents: 21439
diff changeset
    36
  val internal: (morphism -> attribute) -> src
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    37
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    38
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    39
structure Attrib: ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    40
struct
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    41
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    42
type src = Args.src;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    43
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    44
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    45
(** named attributes **)
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    46
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    47
(* theory data *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    48
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    49
structure AttributesData = TheoryDataFun
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    50
(
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    51
  type T = (((src -> attribute) * string) * stamp) NameSpace.table;
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    52
  val empty = NameSpace.empty_table;
6546
995a66249a9b theory data: copy;
wenzelm
parents: 6448
diff changeset
    53
  val copy = I;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    54
  val extend = I;
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    55
  fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    56
    error ("Attempt to merge different versions of attribute " ^ quote dup);
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    57
);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    58
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    59
fun print_attributes thy =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    60
  let
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    61
    val attribs = AttributesData.get thy;
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    62
    fun prt_attr (name, ((_, comment), _)) = Pretty.block
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    63
      [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    64
  in
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    65
    [Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))]
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    66
    |> Pretty.chunks |> Pretty.writeln
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    67
  end;
7611
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7598
diff changeset
    68
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    69
21031
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    70
(* name space *)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    71
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    72
val intern = NameSpace.intern o #1 o AttributesData.get;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    73
val intern_src = Args.map_name o intern;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    74
21031
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    75
val extern = NameSpace.extern o #1 o AttributesData.get o ProofContext.theory_of;
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    76
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    77
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    78
(* pretty printing *)
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    79
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    80
fun pretty_attribs _ [] = []
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    81
  | pretty_attribs ctxt srcs =
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    82
      [Pretty.enclose "[" "]"
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    83
        (Pretty.commas (map (Args.pretty_src ctxt o Args.map_name (extern ctxt)) srcs))];
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    84
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    85
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    86
(* get attributes *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    87
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    88
fun attribute_i thy =
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    89
  let
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    90
    val attrs = #2 (AttributesData.get thy);
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
    91
    fun attr src =
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    92
      let val ((name, _), pos) = Args.dest_src src in
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17221
diff changeset
    93
        (case Symtab.lookup attrs name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
    94
          NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
23937
66e1f24d655d eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents: 23655
diff changeset
    95
        | SOME ((att, _), _) => att src)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    96
      end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    97
  in attr end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    98
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    99
fun attribute thy = attribute_i thy o intern_src thy;
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   100
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   101
17105
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   102
(* attributed declarations *)
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   103
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   104
fun map_specs f = map (apfst (apsnd (map f)));
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   105
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
   106
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   107
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   108
(* crude_closure *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   109
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   110
(*Produce closure without knowing facts in advance! The following
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   111
  works reasonably well for attribute parsers that do not peek at the
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   112
  thm structure.*)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   113
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   114
fun crude_closure ctxt src =
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   115
 (try (fn () => attribute_i (ProofContext.theory_of ctxt) src
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   116
    (Context.Proof ctxt, Drule.asm_rl)) ();
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   117
  Args.closure src);
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   118
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   119
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   120
(* add_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   121
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   122
fun add_attributes raw_attrs thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   123
  let
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   124
    val new_attrs =
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   125
      raw_attrs |> map (fn (name, att, comment) => (name, ((att, comment), stamp ())));
23086
12320f6e2523 tuned Pure/General/name_space.ML
haftmann
parents: 22900
diff changeset
   126
    fun add attrs = NameSpace.extend_table (Sign.naming_of thy) new_attrs attrs
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
   127
      handle Symtab.DUP dup => error ("Duplicate declaration of attributes " ^ quote dup);
16344
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
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   131
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   132
(** attribute parsers **)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   133
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   134
(* tags *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   135
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
   136
fun tag x = Scan.lift (Args.name -- Args.name) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   137
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   138
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   139
(* theorems *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   140
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   141
local
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   142
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   143
val get_thms = Context.cases PureThy.get_thms ProofContext.get_thms;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   144
21698
43a842769765 thms etc.: proper treatment of internal_fact with selection;
wenzelm
parents: 21658
diff changeset
   145
val fact_name = Args.internal_fact >> K "<fact>" || Args.name;
43a842769765 thms etc.: proper treatment of internal_fact with selection;
wenzelm
parents: 21658
diff changeset
   146
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   147
fun gen_thm pick = Scan.depend (fn st =>
24008
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   148
  Args.$$$ "[" |-- Args.attribs (intern (Context.theory_of st)) --| Args.$$$ "]"
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   149
    >> (fn srcs =>
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   150
      let
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   151
        val atts = map (attribute_i (Context.theory_of st)) srcs;
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   152
        val (st', th') = Library.apply atts (st, Drule.dummy_thm);
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   153
      in (st', pick "" [th']) end) ||
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   154
  (Scan.ahead Args.alt_name -- Args.named_fact (get_thms st o Fact)
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   155
    >> (fn (s, fact) => ("", Fact s, fact)) ||
21698
43a842769765 thms etc.: proper treatment of internal_fact with selection;
wenzelm
parents: 21658
diff changeset
   156
  Scan.ahead fact_name -- Args.named_fact (get_thms st o Name) -- Args.thm_sel
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   157
    >> (fn ((name, fact), sel) => (name, NameSelection (name, sel), fact)) ||
21698
43a842769765 thms etc.: proper treatment of internal_fact with selection;
wenzelm
parents: 21658
diff changeset
   158
  Scan.ahead fact_name -- Args.named_fact (get_thms st o Name)
24008
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   159
    >> (fn (name, fact) => (name, Name name, fact)))
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   160
  -- Args.opt_attribs (intern (Context.theory_of st))
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   161
  >> (fn ((name, thmref, fact), srcs) =>
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   162
    let
16498
9d265401fee0 thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   163
      val ths = PureThy.select_thm thmref fact;
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   164
      val atts = map (attribute_i (Context.theory_of st)) srcs;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   165
      val (st', ths') = foldl_map (Library.apply atts) (st, ths);
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   166
    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
   167
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   168
in
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   169
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   170
val thm = gen_thm PureThy.single_thm;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   171
val multi_thm = gen_thm (K I);
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19473
diff changeset
   172
val thms = Scan.repeat multi_thm >> flat;
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   173
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   174
end;
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   175
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   176
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   177
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   178
(** attribute syntax **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   179
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   180
fun syntax scan src (st, th) =
21879
a3efbae45735 switched argument order in *.syntax lifters
haftmann
parents: 21698
diff changeset
   181
  let val (f, st') = Args.syntax "attribute" scan src st
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   182
  in f (st', th) end;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   183
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   184
fun no_args x = syntax (Scan.succeed x);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   185
10034
4bca6b2d2589 tuned args;
wenzelm
parents: 9952
diff changeset
   186
fun add_del_args add del x = syntax
4bca6b2d2589 tuned args;
wenzelm
parents: 9952
diff changeset
   187
  (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
   188
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   189
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   190
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   191
(** basic attributes **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   192
23988
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   193
(* configuration options *)
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   194
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   195
local
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   196
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   197
val equals = Args.$$$ "=";
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   198
24008
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   199
fun scan_value (ConfigOption.Bool _) =
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   200
      (equals -- Args.$$$ "false") >> K (ConfigOption.Bool false) ||
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   201
      (equals -- Args.$$$ "true") >> K (ConfigOption.Bool true) ||
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   202
      (Scan.succeed (ConfigOption.Bool true))
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   203
  | scan_value (ConfigOption.Int _) = (equals |-- Args.int) >> ConfigOption.Int
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   204
  | scan_value (ConfigOption.String _) = (equals |-- Args.name) >> ConfigOption.String;
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   205
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   206
fun scan_config x =
24022
wenzelm
parents: 24008
diff changeset
   207
  ((Args.name >> ConfigOption.the_option) :|-- (fn (config, config_type) =>
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   208
    scan_value config_type >> (fn value =>
24022
wenzelm
parents: 24008
diff changeset
   209
      K (Thm.declaration_attribute (K (ConfigOption.put_generic config value)))))) x;
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   210
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   211
fun scan_att x =
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   212
  (Args.internal_attribute ||
24022
wenzelm
parents: 24008
diff changeset
   213
    (Scan.ahead (scan_config --| Args.terminator) :|--
wenzelm
parents: 24008
diff changeset
   214
      (fn att => Args.named_attribute (K att)))) x;
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   215
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   216
in
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   217
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   218
fun option x = syntax (Scan.lift ((scan_att >> Morphism.form) --| Scan.many Args.not_eof)) x;
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   219
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   220
end;
23988
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   221
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   222
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   223
(* tags *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   224
18799
f137c5e971f5 moved theorem tags from Drule to PureThy;
wenzelm
parents: 18734
diff changeset
   225
fun tagged x = syntax (tag >> PureThy.tag) x;
f137c5e971f5 moved theorem tags from Drule to PureThy;
wenzelm
parents: 18734
diff changeset
   226
fun untagged x = syntax (Scan.lift Args.name >> PureThy.untag) x;
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   227
21439
303ec9e9f74f removed kind attribs;
wenzelm
parents: 21031
diff changeset
   228
fun kind x = syntax (Scan.lift Args.name >> PureThy.kind) x;
20906
63150f3a103d added kind attributes;
wenzelm
parents: 20334
diff changeset
   229
63150f3a103d added kind attributes;
wenzelm
parents: 20334
diff changeset
   230
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   231
(* rule composition *)
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   232
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   233
val COMP_att =
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   234
  syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   235
    >> (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
   236
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   237
val THEN_att =
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   238
  syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   239
    >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B))));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   240
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   241
val OF_att =
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   242
  syntax (thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A)));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   243
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   244
13782
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   245
(* rename_abs *)
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   246
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   247
fun rename_abs src = syntax
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   248
  (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
   249
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   250
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   251
(* unfold / fold definitions *)
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   252
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   253
fun unfolded_syntax rule =
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   254
  syntax (thms >>
18872
020044cf1510 (un)folded: removed '(raw)' option;
wenzelm
parents: 18839
diff changeset
   255
    (fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths)));
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   256
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   257
val unfolded = unfolded_syntax LocalDefs.unfold;
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   258
val folded = unfolded_syntax LocalDefs.fold;
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   259
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   260
8368
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   261
(* rule cases *)
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   262
10528
a0483268262d added "consumes" attribute;
wenzelm
parents: 10151
diff changeset
   263
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
   264
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
   265
fun case_conclusion x =
dd445f5cb28e added case_conclusion attribute;
wenzelm
parents: 18037
diff changeset
   266
  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
   267
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
   268
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   269
20241
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   270
(* rule format *)
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   271
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   272
fun rule_format_att x = syntax (Args.mode "no_asm"
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   273
  >> (fn true => ObjectLogic.rule_format_no_asm | false => ObjectLogic.rule_format)) x;
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   274
20241
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   275
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
   276
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   277
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   278
(* misc rules *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   279
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   280
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
   281
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   282
fun no_vars x = no_args (Thm.rule_attribute (fn ctxt => fn th =>
22568
ed7aa5a350ef renamed Variable.import to import_thms (avoid clash with Alice keywords);
wenzelm
parents: 22115
diff changeset
   283
  let val ((_, [th']), _) = Variable.import_thms true [th] (Context.proof_of ctxt)
20241
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   284
  in th' end)) x;
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   285
22900
f8a7c10e1bd0 moved conversions to structure Conv;
wenzelm
parents: 22846
diff changeset
   286
fun eta_long x =
f8a7c10e1bd0 moved conversions to structure Conv;
wenzelm
parents: 22846
diff changeset
   287
  no_args (Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion))) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   288
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   289
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   290
(* internal attribute *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   291
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   292
fun internal att = Args.src (("Pure.attribute", [Args.mk_attribute att]), Position.none);
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   293
21658
5e31241e1e3c Attrib.internal: morphism;
wenzelm
parents: 21439
diff changeset
   294
fun internal_att x =
22669
62857ad97cca Morphism.transform/form;
wenzelm
parents: 22568
diff changeset
   295
  syntax (Scan.lift Args.internal_attribute >> Morphism.form) x;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   296
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   297
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   298
(* theory setup *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   299
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   300
val _ = Context.add_setup
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18678
diff changeset
   301
 (add_attributes
23988
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   302
   [("option", option, "named configuration option"),
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   303
    ("tagged", tagged, "tagged theorem"),
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   304
    ("untagged", untagged, "untagged theorem"),
21439
303ec9e9f74f removed kind attribs;
wenzelm
parents: 21031
diff changeset
   305
    ("kind", kind, "theorem kind"),
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   306
    ("COMP", COMP_att, "direct composition with rules (no lifting)"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   307
    ("THEN", THEN_att, "resolution with rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   308
    ("OF", OF_att, "rule applied to facts"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   309
    ("rename_abs", rename_abs, "rename bound variables in abstractions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   310
    ("unfolded", unfolded, "unfolded definitions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   311
    ("folded", folded, "folded definitions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   312
    ("standard", standard, "result put into standard form"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   313
    ("elim_format", elim_format, "destruct rule turned into elimination rule format"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   314
    ("no_vars", no_vars, "frozen schematic vars"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   315
    ("eta_long", eta_long, "put theorem into eta long beta normal form"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   316
    ("consumes", consumes, "number of consumed facts"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   317
    ("case_names", case_names, "named rule cases"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   318
    ("case_conclusion", case_conclusion, "named conclusion of rule cases"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   319
    ("params", params, "named rule parameters"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   320
    ("atomize", no_args ObjectLogic.declare_atomize, "declaration of atomize rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   321
    ("rulify", no_args ObjectLogic.declare_rulify, "declaration of rulify rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   322
    ("rule_format", rule_format_att, "result put into standard rule format"),
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   323
    ("defn", add_del_args LocalDefs.defn_add LocalDefs.defn_del,
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   324
      "declaration of definitional transformations"),
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   325
    ("attribute", internal_att, "internal attribute")]);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   326
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   327
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   328
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   329
structure BasicAttrib: BASIC_ATTRIB = Attrib;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   330
open BasicAttrib;