src/Pure/Isar/attrib.ML
author haftmann
Mon, 17 Dec 2007 22:40:14 +0100
changeset 25684 2b3cce7d22b8
parent 24723 2110df1e157d
child 25983 111d2ed164f4
permissions -rw-r--r--
note in target
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
24723
2110df1e157d added eval_thms;
wenzelm
parents: 24713
diff changeset
    22
  val eval_thms: Proof.context -> (thmref * src list) list -> thm list
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
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    30
  val print_configs: Proof.context -> unit
24126
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
    31
  val register_config: Config.value Config.T -> theory -> theory
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    32
  val config_bool: bstring -> bool -> bool Config.T * (theory -> theory)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    33
  val config_int: bstring -> int -> int Config.T * (theory -> theory)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    34
  val config_string: bstring -> string -> string Config.T * (theory -> theory)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    35
  val config_bool_global: bstring -> bool -> bool Config.T * (theory -> theory)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    36
  val config_int_global: bstring -> int -> int Config.T * (theory -> theory)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    37
  val config_string_global: bstring -> string -> string Config.T * (theory -> theory)
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    38
  val thm: Context.generic * Args.T list -> thm * (Context.generic * Args.T list)
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    39
  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
    40
  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
    41
  val syntax: (Context.generic * Args.T list ->
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    42
    attribute * (Context.generic * Args.T list)) -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    43
  val no_args: attribute -> src -> attribute
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    44
  val add_del_args: attribute -> attribute -> src -> attribute
21658
5e31241e1e3c Attrib.internal: morphism;
wenzelm
parents: 21439
diff changeset
    45
  val internal: (morphism -> attribute) -> src
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    46
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    47
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    48
structure Attrib: ATTRIB =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    49
struct
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    50
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    51
type src = Args.src;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    52
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    53
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    54
(** named attributes **)
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
    55
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    56
(* theory data *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    57
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    58
structure Attributes = TheoryDataFun
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    59
(
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    60
  type T = (((src -> attribute) * string) * stamp) NameSpace.table;
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
    61
  val empty = NameSpace.empty_table;
6546
995a66249a9b theory data: copy;
wenzelm
parents: 6448
diff changeset
    62
  val copy = I;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16344
diff changeset
    63
  val extend = I;
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
    64
  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
    65
    error ("Attempt to merge different versions of attribute " ^ quote dup);
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    66
);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    67
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    68
fun print_attributes thy =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    69
  let
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    70
    val attribs = Attributes.get thy;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    71
    fun prt_attr (name, ((_, comment), _)) = Pretty.block
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    72
      [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    73
  in
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    74
    [Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))]
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    75
    |> Pretty.chunks |> Pretty.writeln
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22669
diff changeset
    76
  end;
7611
5b5aba10c8f6 help: unkown theory context;
wenzelm
parents: 7598
diff changeset
    77
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    78
21031
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    79
(* name space *)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    80
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    81
val intern = NameSpace.intern o #1 o Attributes.get;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    82
val intern_src = Args.map_name o intern;
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    83
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    84
val extern = NameSpace.extern o #1 o Attributes.get o ProofContext.theory_of;
21031
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    85
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    86
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    87
(* pretty printing *)
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    88
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    89
fun pretty_attribs _ [] = []
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    90
  | pretty_attribs ctxt srcs =
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    91
      [Pretty.enclose "[" "]"
a56e6d1e56a3 added pretty_attribs (from attrib.ML);
wenzelm
parents: 20906
diff changeset
    92
        (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
    93
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
    94
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    95
(* get attributes *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    96
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    97
fun attribute_i thy =
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
    98
  let
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
    99
    val attrs = #2 (Attributes.get thy);
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   100
    fun attr src =
16344
a52fe1277902 attribs: NameSpace.table;
wenzelm
parents: 16141
diff changeset
   101
      let val ((name, _), pos) = Args.dest_src src in
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17221
diff changeset
   102
        (case Symtab.lookup attrs name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   103
          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
   104
        | SOME ((att, _), _) => att src)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   105
      end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   106
  in attr end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   107
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   108
fun attribute thy = attribute_i thy o intern_src thy;
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   109
24723
2110df1e157d added eval_thms;
wenzelm
parents: 24713
diff changeset
   110
fun eval_thms ctxt args = ProofContext.note_thmss Thm.theoremK
2110df1e157d added eval_thms;
wenzelm
parents: 24713
diff changeset
   111
  [(("", []), map (apsnd (map (attribute (ProofContext.theory_of ctxt)))) args)] ctxt
2110df1e157d added eval_thms;
wenzelm
parents: 24713
diff changeset
   112
  |> (flat o map snd o fst);
2110df1e157d added eval_thms;
wenzelm
parents: 24713
diff changeset
   113
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   114
17105
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   115
(* attributed declarations *)
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   116
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   117
fun map_specs f = map (apfst (apsnd (map f)));
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   118
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
   119
1b07a176a880 added map_specs/facts operators (from locale.ML);
wenzelm
parents: 16934
diff changeset
   120
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   121
(* crude_closure *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   122
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   123
(*Produce closure without knowing facts in advance! The following
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   124
  works reasonably well for attribute parsers that do not peek at the
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   125
  thm structure.*)
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   126
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   127
fun crude_closure ctxt src =
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   128
 (try (fn () => attribute_i (ProofContext.theory_of ctxt) src
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   129
    (Context.Proof ctxt, Drule.asm_rl)) ();
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   130
  Args.closure src);
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   131
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   132
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   133
(* add_attributes *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   134
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   135
fun add_attributes raw_attrs thy =
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   136
  let
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   137
    val new_attrs =
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   138
      raw_attrs |> map (fn (name, att, comment) => (name, ((att, comment), stamp ())));
23086
12320f6e2523 tuned Pure/General/name_space.ML
haftmann
parents: 22900
diff changeset
   139
    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
   140
      handle Symtab.DUP dup => error ("Duplicate declaration of attributes " ^ quote dup);
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   141
  in Attributes.map add thy end;
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   142
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   143
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   144
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   145
(** attribute parsers **)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   146
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   147
(* tags *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   148
23655
d2d1138e0ddc replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
wenzelm
parents: 23577
diff changeset
   149
fun tag x = Scan.lift (Args.name -- Args.name) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   150
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   151
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   152
(* theorems *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   153
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   154
local
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   155
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   156
val get_thms = Context.cases PureThy.get_thms ProofContext.get_thms;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   157
21698
43a842769765 thms etc.: proper treatment of internal_fact with selection;
wenzelm
parents: 21658
diff changeset
   158
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
   159
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   160
fun gen_thm pick = Scan.depend (fn st =>
24008
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   161
  Args.$$$ "[" |-- Args.attribs (intern (Context.theory_of st)) --| Args.$$$ "]"
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   162
    >> (fn srcs =>
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   163
      let
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   164
        val atts = map (attribute_i (Context.theory_of st)) srcs;
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   165
        val (st', th') = Library.apply atts (st, Drule.dummy_thm);
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   166
      in (st', pick "" [th']) end) ||
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   167
  (Scan.ahead Args.alt_name -- Args.named_fact (get_thms st o Fact)
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   168
    >> (fn (s, fact) => ("", Fact s, fact)) ||
21698
43a842769765 thms etc.: proper treatment of internal_fact with selection;
wenzelm
parents: 21658
diff changeset
   169
  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
   170
    >> (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
   171
  Scan.ahead fact_name -- Args.named_fact (get_thms st o Name)
24008
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   172
    >> (fn (name, fact) => (name, Name name, fact)))
63ff2445ce1e renamed Config to ConfigOption;
wenzelm
parents: 24003
diff changeset
   173
  -- Args.opt_attribs (intern (Context.theory_of st))
18037
1095d2213b9d syntax for literal facts;
wenzelm
parents: 17756
diff changeset
   174
  >> (fn ((name, thmref, fact), srcs) =>
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   175
    let
16498
9d265401fee0 thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   176
      val ths = PureThy.select_thm thmref fact;
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   177
      val atts = map (attribute_i (Context.theory_of st)) srcs;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   178
      val (st', ths') = foldl_map (Library.apply atts) (st, ths);
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   179
    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
   180
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   181
in
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   182
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   183
val thm = gen_thm PureThy.single_thm;
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   184
val multi_thm = gen_thm (K I);
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19473
diff changeset
   185
val thms = Scan.repeat multi_thm >> flat;
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   186
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   187
end;
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   188
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   189
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   190
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   191
(** attribute syntax **)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   192
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   193
fun syntax scan src (st, th) =
21879
a3efbae45735 switched argument order in *.syntax lifters
haftmann
parents: 21698
diff changeset
   194
  let val (f, st') = Args.syntax "attribute" scan src st
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   195
  in f (st', th) end;
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   196
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   197
fun no_args x = syntax (Scan.succeed x);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   198
10034
4bca6b2d2589 tuned args;
wenzelm
parents: 9952
diff changeset
   199
fun add_del_args add del x = syntax
4bca6b2d2589 tuned args;
wenzelm
parents: 9952
diff changeset
   200
  (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
   201
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   202
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   203
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   204
(** configuration options **)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   205
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   206
(* naming *)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   207
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   208
structure Configs = TheoryDataFun
24713
8b3b6d09ef40 tuned functor application;
wenzelm
parents: 24238
diff changeset
   209
(
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   210
  type T = Config.value Config.T Symtab.table;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   211
  val empty = Symtab.empty;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   212
  val copy = I;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   213
  val extend = I;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   214
  fun merge _ = Symtab.merge (K true);
24713
8b3b6d09ef40 tuned functor application;
wenzelm
parents: 24238
diff changeset
   215
);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   216
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   217
fun print_configs ctxt =
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   218
  let
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   219
    val thy = ProofContext.theory_of ctxt;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   220
    fun prt (name, config) =
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   221
      let val value = Config.get ctxt config in
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   222
        Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1,
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   223
          Pretty.str (Config.print_value value)]
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   224
      end;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   225
    val configs = NameSpace.extern_table (#1 (Attributes.get thy), Configs.get thy);
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   226
  in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   227
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   228
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   229
(* concrete syntax *)
23988
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   230
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   231
local
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   232
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   233
val equals = Args.$$$ "=";
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   234
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   235
fun scan_value (Config.Bool _) =
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   236
      equals -- Args.$$$ "false" >> K (Config.Bool false) ||
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   237
      equals -- Args.$$$ "true" >> K (Config.Bool true) ||
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   238
      Scan.succeed (Config.Bool true)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   239
  | scan_value (Config.Int _) = equals |-- Args.int >> Config.Int
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   240
  | scan_value (Config.String _) = equals |-- Args.name >> Config.String;
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   241
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   242
fun scan_config thy config =
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   243
  let val config_type = Config.get_thy thy config
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   244
  in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end;
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   245
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   246
in
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   247
24126
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   248
fun register_config config thy =
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   249
  let
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   250
    val bname = Config.name_of config;
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   251
    val name = Sign.full_name thy bname;
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   252
  in
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   253
    thy
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   254
    |> add_attributes [(bname, syntax (Scan.lift (scan_config thy config) >> Morphism.form),
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   255
      "configuration option")]
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   256
    |> Configs.map (Symtab.update (name, config))
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   257
  end;
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   258
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   259
fun declare_config make coerce global name default =
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   260
  let
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   261
    val config_value = Config.declare global name (make default);
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   262
    val config = coerce config_value;
24126
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   263
  in (config, register_config config_value) end;
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   264
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   265
val config_bool   = declare_config Config.Bool Config.bool false;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   266
val config_int    = declare_config Config.Int Config.int false;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   267
val config_string = declare_config Config.String Config.string false;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   268
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   269
val config_bool_global   = declare_config Config.Bool Config.bool true;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   270
val config_int_global    = declare_config Config.Int Config.int true;
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   271
val config_string_global = declare_config Config.String Config.string true;
24003
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   272
da76d7e6435c attribute "option": more elaborate syntax (with value parsing);
wenzelm
parents: 23988
diff changeset
   273
end;
23988
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   274
aa46577f4f44 added attribute "option" for setting configuration options;
wenzelm
parents: 23937
diff changeset
   275
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   276
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   277
(** basic attributes **)
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   278
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   279
(* tags *)
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   280
18799
f137c5e971f5 moved theorem tags from Drule to PureThy;
wenzelm
parents: 18734
diff changeset
   281
fun tagged x = syntax (tag >> PureThy.tag) x;
f137c5e971f5 moved theorem tags from Drule to PureThy;
wenzelm
parents: 18734
diff changeset
   282
fun untagged x = syntax (Scan.lift Args.name >> PureThy.untag) x;
6772
111845fce1b7 added COMP attribute;
wenzelm
parents: 6546
diff changeset
   283
21439
303ec9e9f74f removed kind attribs;
wenzelm
parents: 21031
diff changeset
   284
fun kind x = syntax (Scan.lift Args.name >> PureThy.kind) x;
20906
63150f3a103d added kind attributes;
wenzelm
parents: 20334
diff changeset
   285
63150f3a103d added kind attributes;
wenzelm
parents: 20334
diff changeset
   286
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   287
(* rule composition *)
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   288
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   289
val COMP_att =
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   290
  syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   291
    >> (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
   292
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   293
val THEN_att =
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   294
  syntax (Scan.lift (Scan.optional (Args.bracks Args.nat) 1) -- thm
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   295
    >> (fn (i, B) => Thm.rule_attribute (fn _ => fn A => A RSN (i, B))));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   296
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   297
val OF_att =
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   298
  syntax (thms >> (fn Bs => Thm.rule_attribute (fn _ => fn A => Bs MRS A)));
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   299
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   300
13782
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   301
(* rename_abs *)
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   302
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   303
fun rename_abs src = syntax
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   304
  (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
   305
44de406a7273 Added rename_abs attribute for renaming bound variables.
berghofe
parents: 13414
diff changeset
   306
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   307
(* unfold / fold definitions *)
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   308
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   309
fun unfolded_syntax rule =
18998
10c251f29847 Context.generic is canonical state of parsers;
wenzelm
parents: 18977
diff changeset
   310
  syntax (thms >>
18872
020044cf1510 (un)folded: removed '(raw)' option;
wenzelm
parents: 18839
diff changeset
   311
    (fn ths => Thm.rule_attribute (fn context => rule (Context.proof_of context) ths)));
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   312
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   313
val unfolded = unfolded_syntax LocalDefs.unfold;
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   314
val folded = unfolded_syntax LocalDefs.fold;
7598
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   315
af320257c902 unfold / fold defs;
wenzelm
parents: 7367
diff changeset
   316
8368
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   317
(* rule cases *)
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   318
10528
a0483268262d added "consumes" attribute;
wenzelm
parents: 10151
diff changeset
   319
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
   320
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
   321
fun case_conclusion x =
dd445f5cb28e added case_conclusion attribute;
wenzelm
parents: 18037
diff changeset
   322
  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
   323
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
   324
bdc3ee0d8cb6 added 'case_names' and 'params';
wenzelm
parents: 8282
diff changeset
   325
20241
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   326
(* rule format *)
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   327
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   328
fun rule_format_att x = syntax (Args.mode "no_asm"
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   329
  >> (fn true => ObjectLogic.rule_format_no_asm | false => ObjectLogic.rule_format)) x;
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   330
20241
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   331
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
   332
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11763
diff changeset
   333
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   334
(* misc rules *)
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   335
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   336
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
   337
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   338
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
   339
  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
   340
  in th' end)) x;
a571d044891e no_vars: based on Variable.import;
wenzelm
parents: 20029
diff changeset
   341
22900
f8a7c10e1bd0 moved conversions to structure Conv;
wenzelm
parents: 22846
diff changeset
   342
fun eta_long x =
f8a7c10e1bd0 moved conversions to structure Conv;
wenzelm
parents: 22846
diff changeset
   343
  no_args (Thm.rule_attribute (K (Conv.fconv_rule Drule.eta_long_conversion))) x;
5879
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   344
18b8f048d93a several args parsers;
wenzelm
parents: 5823
diff changeset
   345
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   346
(* internal attribute *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   347
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   348
fun internal att = Args.src (("Pure.attribute", [Args.mk_attribute att]), Position.none);
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   349
21658
5e31241e1e3c Attrib.internal: morphism;
wenzelm
parents: 21439
diff changeset
   350
fun internal_att x =
22669
62857ad97cca Morphism.transform/form;
wenzelm
parents: 22568
diff changeset
   351
  syntax (Scan.lift Args.internal_attribute >> Morphism.form) x;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   352
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15596
diff changeset
   353
24238
ae70f95e31de new attribute [rotated]
kleing
parents: 24178
diff changeset
   354
(* attribute rotated *)
ae70f95e31de new attribute [rotated]
kleing
parents: 24178
diff changeset
   355
ae70f95e31de new attribute [rotated]
kleing
parents: 24178
diff changeset
   356
fun rotated_att x = 
ae70f95e31de new attribute [rotated]
kleing
parents: 24178
diff changeset
   357
  syntax (Scan.lift (Scan.optional Args.int 1) >> 
ae70f95e31de new attribute [rotated]
kleing
parents: 24178
diff changeset
   358
                    (fn n => Thm.rule_attribute (fn _ => rotate_prems n))) x
ae70f95e31de new attribute [rotated]
kleing
parents: 24178
diff changeset
   359
ae70f95e31de new attribute [rotated]
kleing
parents: 24178
diff changeset
   360
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   361
(* theory setup *)
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   362
18636
cb068cfdcac8 added rule, declaration;
wenzelm
parents: 18418
diff changeset
   363
val _ = Context.add_setup
24178
4ff1dc2aa18d turned Unify flags into configuration options (global only);
wenzelm
parents: 24126
diff changeset
   364
 (register_config Unify.trace_bound_value #>
4ff1dc2aa18d turned Unify flags into configuration options (global only);
wenzelm
parents: 24126
diff changeset
   365
  register_config Unify.search_bound_value #>
4ff1dc2aa18d turned Unify flags into configuration options (global only);
wenzelm
parents: 24126
diff changeset
   366
  register_config Unify.trace_simp_value #>
4ff1dc2aa18d turned Unify flags into configuration options (global only);
wenzelm
parents: 24126
diff changeset
   367
  register_config Unify.trace_types_value #>
4ff1dc2aa18d turned Unify flags into configuration options (global only);
wenzelm
parents: 24126
diff changeset
   368
  register_config MetaSimplifier.simp_depth_limit_value #>
24126
913e1fa904fb turned simp_depth_limit into configuration option;
wenzelm
parents: 24110
diff changeset
   369
  add_attributes
24110
4ab3084e311c tuned config options: eliminated separate attribute "option";
wenzelm
parents: 24030
diff changeset
   370
   [("tagged", tagged, "tagged theorem"),
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   371
    ("untagged", untagged, "untagged theorem"),
21439
303ec9e9f74f removed kind attribs;
wenzelm
parents: 21031
diff changeset
   372
    ("kind", kind, "theorem kind"),
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   373
    ("COMP", COMP_att, "direct composition with rules (no lifting)"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   374
    ("THEN", THEN_att, "resolution with rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   375
    ("OF", OF_att, "rule applied to facts"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   376
    ("rename_abs", rename_abs, "rename bound variables in abstractions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   377
    ("unfolded", unfolded, "unfolded definitions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   378
    ("folded", folded, "folded definitions"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   379
    ("standard", standard, "result put into standard form"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   380
    ("elim_format", elim_format, "destruct rule turned into elimination rule format"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   381
    ("no_vars", no_vars, "frozen schematic vars"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   382
    ("eta_long", eta_long, "put theorem into eta long beta normal form"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   383
    ("consumes", consumes, "number of consumed facts"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   384
    ("case_names", case_names, "named rule cases"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   385
    ("case_conclusion", case_conclusion, "named conclusion of rule cases"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   386
    ("params", params, "named rule parameters"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   387
    ("atomize", no_args ObjectLogic.declare_atomize, "declaration of atomize rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   388
    ("rulify", no_args ObjectLogic.declare_rulify, "declaration of rulify rule"),
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   389
    ("rule_format", rule_format_att, "result put into standard rule format"),
24238
ae70f95e31de new attribute [rotated]
kleing
parents: 24178
diff changeset
   390
    ("rotated", rotated_att, "rotated theorem premises"),
18839
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   391
    ("defn", add_del_args LocalDefs.defn_add LocalDefs.defn_del,
86a59812b03b added 'defn' attribute;
wenzelm
parents: 18821
diff changeset
   392
      "declaration of definitional transformations"),
18734
f5ea6b0d3501 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   393
    ("attribute", internal_att, "internal attribute")]);
5823
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   394
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   395
end;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   396
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   397
structure BasicAttrib: BASIC_ATTRIB = Attrib;
ee7c198a2154 Symbolic theorem attributes.
wenzelm
parents:
diff changeset
   398
open BasicAttrib;