src/Pure/Isar/context_rules.ML
author wenzelm
Fri, 11 Jul 2025 11:52:43 +0200
changeset 82839 60ec2b2dc55a
parent 82836 68a0219861b7
child 82847 d8ecaff223ff
permissions -rw-r--r--
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/context_rules.ML
82816
wenzelm
parents: 82812
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
wenzelm
parents: 82812
diff changeset
     3
    Author:     Makarius
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
     4
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
     5
Declarations of intro/elim/dest rules in Pure (see also
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
     6
Provers/classical.ML for a more specialized version of the same idea).
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
     7
*)
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
     8
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
     9
signature CONTEXT_RULES =
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    10
sig
82812
ea8d633fd4a8 just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents: 82808
diff changeset
    11
  val netpair_bang: Proof.context -> Bires.netpair
ea8d633fd4a8 just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents: 82808
diff changeset
    12
  val netpair: Proof.context -> Bires.netpair
ea8d633fd4a8 just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents: 82808
diff changeset
    13
  val find_rules_netpair: Proof.context -> bool -> thm list -> term -> Bires.netpair -> thm list
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
    14
  val find_rules: Proof.context -> bool -> thm list -> term -> thm list list
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 20289
diff changeset
    15
  val print_rules: Proof.context -> unit
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51584
diff changeset
    16
  val addSWrapper: (Proof.context -> (int -> tactic) -> int -> tactic) -> theory -> theory
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51584
diff changeset
    17
  val addWrapper: (Proof.context -> (int -> tactic) -> int -> tactic) -> theory -> theory
20289
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 19482
diff changeset
    18
  val Swrap: Proof.context -> (int -> tactic) -> int -> tactic
ba7a7c56bed5 normalized Proof.context/method type aliases;
wenzelm
parents: 19482
diff changeset
    19
  val wrap: Proof.context -> (int -> tactic) -> int -> tactic
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    20
  val intro_bang: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    21
  val elim_bang: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    22
  val dest_bang: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    23
  val intro: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    24
  val elim: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    25
  val dest: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    26
  val intro_query: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    27
  val elim_query: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    28
  val dest_query: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    29
  val rule_del: attribute
30529
23d1892f8015 simplified attribute setup;
wenzelm
parents: 29606
diff changeset
    30
  val add: (int option -> attribute) -> (int option -> attribute) -> (int option -> attribute) ->
23d1892f8015 simplified attribute setup;
wenzelm
parents: 29606
diff changeset
    31
    attribute context_parser
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    32
end;
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    33
33369
470a7b233ee5 modernized structure Context_Rules;
wenzelm
parents: 32784
diff changeset
    34
structure Context_Rules: CONTEXT_RULES =
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    35
struct
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    36
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    37
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    38
(** rule declaration contexts **)
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    39
82839
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    40
(* rule declarations *)
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    41
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    42
type decl = thm Bires.decl;
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    43
type decls = thm Bires.decls;
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    44
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    45
fun init_decl kind opt_weight th : decl =
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    46
  let val weight = opt_weight |> \<^if_none>\<open>Bires.subgoals_of (Bires.kind_rule kind th)\<close>;
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    47
  in {kind = kind, tag = Bires.weight_tag weight, info = th} end;
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    48
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    49
fun insert_decl ({kind, tag, info = th}: decl) =
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    50
  Bires.insert_tagged_rule (tag, Bires.kind_rule kind th);
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    51
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    52
fun remove_decl ({info = th, ...}: decl) =
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    53
  Bires.delete_tagged_rule (false, th) o Bires.delete_tagged_rule (true, th);
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    54
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    55
18637
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
    56
(* context data *)
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    57
33370
69531a7b55b6 tuned signature;
wenzelm
parents: 33369
diff changeset
    58
datatype rules = Rules of
82839
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    59
 {decls: decls,
82812
ea8d633fd4a8 just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents: 82808
diff changeset
    60
  netpairs: Bires.netpair list,
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51584
diff changeset
    61
  wrappers:
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51584
diff changeset
    62
    ((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list *
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51584
diff changeset
    63
    ((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list};
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    64
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    65
fun make_rules decls netpairs wrappers =
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    66
  Rules {decls = decls, netpairs = netpairs, wrappers = wrappers};
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    67
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    68
fun add_rule kind opt_weight th (rules as Rules {decls, netpairs, wrappers}) =
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
    69
  let
82839
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    70
    val th' = Thm.trim_context th;
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    71
    val decl' = init_decl kind opt_weight th';
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    72
  in
82839
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    73
    (case Bires.extend_decls (th', decl') decls of
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    74
      (NONE, _) => rules
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    75
    | (SOME new_decl, decls') =>
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    76
        let val netpairs' = Bires.kind_map kind (insert_decl new_decl) netpairs
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    77
        in make_rules decls' netpairs' wrappers end)
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    78
  end;
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    79
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    80
fun del_rule0 th (rules as Rules {decls, netpairs, wrappers}) =
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    81
  (case Bires.remove_decls th decls of
82839
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    82
    ([], _) => rules
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    83
  | (old_decls, decls') =>
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    84
      let
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    85
        val netpairs' =
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    86
          fold (fn decl as {kind, ...} => Bires.kind_map kind (remove_decl decl))
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
    87
            old_decls netpairs;
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    88
      in make_rules decls' netpairs' wrappers end);
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    89
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 39557
diff changeset
    90
fun del_rule th = del_rule0 th o del_rule0 (Tactic.make_elim th);
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 39557
diff changeset
    91
82816
wenzelm
parents: 82812
diff changeset
    92
structure Data = Generic_Data
18637
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
    93
(
33370
69531a7b55b6 tuned signature;
wenzelm
parents: 33369
diff changeset
    94
  type T = rules;
82828
05d2b8b675da clarified signature;
wenzelm
parents: 82826
diff changeset
    95
  val empty = make_rules Bires.empty_decls (Bires.kind_values Bires.empty_netpair) ([], []);
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33370
diff changeset
    96
  fun merge
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    97
    (Rules {decls = decls1, netpairs = netpairs1, wrappers = (ws1, ws1')},
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
    98
     Rules {decls = decls2, netpairs = _, wrappers = (ws2, ws2')}) =
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
    99
    let
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
   100
      val (new_rules, decls) = Bires.merge_decls (decls1, decls2);
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
   101
      val netpairs =
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
   102
        netpairs1 |> map_index (uncurry (fn i =>
82839
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
   103
          new_rules |> fold (fn decl =>
60ec2b2dc55a clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents: 82836
diff changeset
   104
            if Bires.kind_index (#kind decl) = i then insert_decl decl else I)))
18637
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
   105
      val wrappers =
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
   106
       (Library.merge (eq_snd (op =)) (ws1, ws2),
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
   107
        Library.merge (eq_snd (op =)) (ws1', ws2'));
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
   108
    in make_rules decls netpairs wrappers end;
18637
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
   109
);
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   110
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22360
diff changeset
   111
fun print_rules ctxt =
82836
68a0219861b7 clarified modules;
wenzelm
parents: 82829
diff changeset
   112
  let val Rules {decls, ...} = Data.get (Context.Proof ctxt)
68a0219861b7 clarified modules;
wenzelm
parents: 82829
diff changeset
   113
  in Pretty.writeln (Pretty.chunks (Bires.pretty_decls ctxt Bires.kind_domain decls)) end;
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   114
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   115
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   116
(* access data *)
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   117
82816
wenzelm
parents: 82812
diff changeset
   118
fun netpairs ctxt = let val Rules {netpairs, ...} = Data.get (Context.Proof ctxt) in netpairs end;
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   119
val netpair_bang = hd o netpairs;
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   120
val netpair = hd o tl o netpairs;
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   121
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   122
12399
wenzelm
parents: 12380
diff changeset
   123
(* retrieving rules *)
wenzelm
parents: 12380
diff changeset
   124
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   125
local
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   126
82812
ea8d633fd4a8 just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents: 82808
diff changeset
   127
fun order weighted =
ea8d633fd4a8 just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents: 82808
diff changeset
   128
  make_order_list (Bires.weighted_tag_ord weighted) NONE;
ea8d633fd4a8 just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents: 82808
diff changeset
   129
12399
wenzelm
parents: 12380
diff changeset
   130
fun may_unify weighted t net =
82812
ea8d633fd4a8 just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents: 82808
diff changeset
   131
  map snd (order weighted (Net.unify_term net t));
12399
wenzelm
parents: 12380
diff changeset
   132
wenzelm
parents: 12380
diff changeset
   133
fun find_erules _ [] = K []
12805
3be853cf19cf Thm.prop_of;
wenzelm
parents: 12412
diff changeset
   134
  | find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact));
16424
18a07ad8fea8 accomodate change of TheoryDataFun;
wenzelm
parents: 15973
diff changeset
   135
12399
wenzelm
parents: 12380
diff changeset
   136
fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal);
12380
3402d300f5ef export low-level addXXs;
wenzelm
parents: 12350
diff changeset
   137
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   138
in
12380
3402d300f5ef export low-level addXXs;
wenzelm
parents: 12350
diff changeset
   139
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   140
fun find_rules_netpair ctxt weighted facts goal (inet, enet) =
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   141
  find_erules weighted facts enet @ find_irules weighted goal inet
67649
1e1782c1aedf tuned signature;
wenzelm
parents: 67626
diff changeset
   142
  |> map (Thm.transfer' ctxt);
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   143
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   144
fun find_rules ctxt weighted facts goal =
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   145
  map (find_rules_netpair ctxt weighted facts goal) (netpairs ctxt);
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   146
0d401f874942 trim context for persistent storage;
wenzelm
parents: 61037
diff changeset
   147
end;
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   148
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   149
12399
wenzelm
parents: 12380
diff changeset
   150
(* wrappers *)
wenzelm
parents: 12380
diff changeset
   151
18667
wenzelm
parents: 18637
diff changeset
   152
fun gen_add_wrapper upd w =
82826
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
   153
  Context.theory_map (Data.map (fn Rules {decls, netpairs, wrappers} =>
f5fd9b41188a efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents: 82818
diff changeset
   154
    make_rules decls netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers)));
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   155
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   156
val addSWrapper = gen_add_wrapper Library.apfst;
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   157
val addWrapper = gen_add_wrapper Library.apsnd;
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   158
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   159
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   160
fun gen_wrap which ctxt =
82816
wenzelm
parents: 82812
diff changeset
   161
  let val Rules {wrappers, ...} = Data.get (Context.Proof ctxt)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51584
diff changeset
   162
  in fold_rev (fn (w, _) => w ctxt) (which wrappers) end;
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   163
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   164
val Swrap = gen_wrap #1;
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   165
val wrap = gen_wrap #2;
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   166
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   167
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   168
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   169
(** attributes **)
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   170
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   171
(* add and del rules *)
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   172
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 39557
diff changeset
   173
82816
wenzelm
parents: 82812
diff changeset
   174
val rule_del = Thm.declaration_attribute (Data.map o del_rule);
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   175
18637
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
   176
fun rule_add k view opt_w =
82816
wenzelm
parents: 82812
diff changeset
   177
  Thm.declaration_attribute (fn th => Data.map (add_rule k opt_w (view th) o del_rule th));
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   178
82817
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   179
val intro_bang  = rule_add Bires.intro_bang_kind I;
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   180
val elim_bang   = rule_add Bires.elim_bang_kind I;
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   181
val dest_bang   = rule_add Bires.elim_bang_kind Tactic.make_elim;
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   182
val intro       = rule_add Bires.intro_kind I;
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   183
val elim        = rule_add Bires.elim_kind I;
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   184
val dest        = rule_add Bires.elim_kind Tactic.make_elim;
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   185
val intro_query = rule_add Bires.intro_query_kind I;
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   186
val elim_query  = rule_add Bires.elim_query_kind I;
be1fb22d9e2a clarified signature: more explicit type Bires.kind;
wenzelm
parents: 82816
diff changeset
   187
val dest_query  = rule_add Bires.elim_query_kind Tactic.make_elim;
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   188
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 51798
diff changeset
   189
val _ = Theory.setup
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 51798
diff changeset
   190
  (snd o Global_Theory.add_thms [((Binding.empty, Drule.equal_intr_rule), [intro_query NONE])]);
18637
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
   191
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   192
18637
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
   193
(* concrete syntax *)
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
   194
30529
23d1892f8015 simplified attribute setup;
wenzelm
parents: 29606
diff changeset
   195
fun add a b c x =
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 26463
diff changeset
   196
  (Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) --
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 33519
diff changeset
   197
    Scan.option Parse.nat) >> (fn (f, n) => f n)) x;
18637
33a6f6caa617 generic data and attributes;
wenzelm
parents: 18418
diff changeset
   198
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 51798
diff changeset
   199
val _ = Theory.setup
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
   200
 (Attrib.setup \<^binding>\<open>intro\<close> (add intro_bang intro intro_query)
82818
c6b3f0ee0a69 tuned messages;
wenzelm
parents: 82817
diff changeset
   201
    "declaration of Pure introduction rule" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
   202
  Attrib.setup \<^binding>\<open>elim\<close> (add elim_bang elim elim_query)
82818
c6b3f0ee0a69 tuned messages;
wenzelm
parents: 82817
diff changeset
   203
    "declaration of Pure elimination rule" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
   204
  Attrib.setup \<^binding>\<open>dest\<close> (add dest_bang dest dest_query)
82818
c6b3f0ee0a69 tuned messages;
wenzelm
parents: 82817
diff changeset
   205
    "declaration of Pure destruction rule" #>
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 61268
diff changeset
   206
  Attrib.setup \<^binding>\<open>rule\<close> (Scan.lift Args.del >> K rule_del)
82818
c6b3f0ee0a69 tuned messages;
wenzelm
parents: 82817
diff changeset
   207
    "remove declaration of Pure intro/elim/dest rule");
12350
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   208
5fad0e7129c3 renamed rule_context.ML to context_rules.ML;
wenzelm
parents:
diff changeset
   209
end;