1(*  Title:      Pure/Isar/context_rules.ML
2    Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
3
4Declarations of intro/elim/dest rules in Pure (see also
5Provers/classical.ML for a more specialized version of the same idea).
6*)
7
8signature CONTEXT_RULES =
9sig
10  type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net
11  val netpair_bang: Proof.context -> netpair
12  val netpair: Proof.context -> netpair
13  val orderlist: ((int * int) * 'a) list -> 'a list
14  val find_rules_netpair: Proof.context -> bool -> thm list -> term -> netpair -> thm list
15  val find_rules: Proof.context -> bool -> thm list -> term -> thm list list
16  val print_rules: Proof.context -> unit
17  val addSWrapper: (Proof.context -> (int -> tactic) -> int -> tactic) -> theory -> theory
18  val addWrapper: (Proof.context -> (int -> tactic) -> int -> tactic) -> theory -> theory
19  val Swrap: Proof.context -> (int -> tactic) -> int -> tactic
20  val wrap: Proof.context -> (int -> tactic) -> int -> tactic
21  val intro_bang: int option -> attribute
22  val elim_bang: int option -> attribute
23  val dest_bang: int option -> attribute
24  val intro: int option -> attribute
25  val elim: int option -> attribute
26  val dest: int option -> attribute
27  val intro_query: int option -> attribute
28  val elim_query: int option -> attribute
29  val dest_query: int option -> attribute
30  val rule_del: attribute
31  val add: (int option -> attribute) -> (int option -> attribute) -> (int option -> attribute) ->
32    attribute context_parser
33end;
34
35structure Context_Rules: CONTEXT_RULES =
36struct
37
38
39(** rule declaration contexts **)
40
41(* rule kinds *)
42
43val intro_bangK = (0, false);
44val elim_bangK = (0, true);
45val introK = (1, false);
46val elimK = (1, true);
47val intro_queryK = (2, false);
48val elim_queryK = (2, true);
49
50val kind_names =
51 [(intro_bangK, "safe introduction rules (intro!)"),
52  (elim_bangK, "safe elimination rules (elim!)"),
53  (introK, "introduction rules (intro)"),
54  (elimK, "elimination rules (elim)"),
55  (intro_queryK, "extra introduction rules (intro?)"),
56  (elim_queryK, "extra elimination rules (elim?)")];
57
58val rule_kinds = map #1 kind_names;
59val rule_indexes = distinct (op =) (map #1 rule_kinds);
60
61
62(* context data *)
63
64type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net;
65val empty_netpairs: netpair list = replicate (length rule_indexes) (Net.empty, Net.empty);
66
67datatype rules = Rules of
68 {next: int,
69  rules: (int * ((int * bool) * thm)) list,
70  netpairs: netpair list,
71  wrappers:
72    ((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list *
73    ((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list};
74
75fun make_rules next rules netpairs wrappers =
76  Rules {next = next, rules = rules, netpairs = netpairs, wrappers = wrappers};
77
78fun add_rule (i, b) opt_w th (Rules {next, rules, netpairs, wrappers}) =
79  let
80    val w = (case opt_w of SOME w => w | NONE => Tactic.subgoals_of_brl (b, th));
81    val th' = Thm.trim_context th;
82  in
83    make_rules (next - 1) ((w, ((i, b), th')) :: rules)
84      (nth_map i (Tactic.insert_tagged_brl ((w, next), (b, th'))) netpairs) wrappers
85  end;
86
87fun del_rule0 th (rs as Rules {next, rules, netpairs, wrappers}) =
88  let
89    fun eq_th (_, (_, th')) = Thm.eq_thm_prop (th, th');
90    fun del b netpair = Tactic.delete_tagged_brl (b, th) netpair handle Net.DELETE => netpair;
91  in
92    if not (exists eq_th rules) then rs
93    else make_rules next (filter_out eq_th rules) (map (del false o del true) netpairs) wrappers
94  end;
95
96fun del_rule th = del_rule0 th o del_rule0 (Tactic.make_elim th);
97
98structure Rules = Generic_Data
99(
100  type T = rules;
101  val empty = make_rules ~1 [] empty_netpairs ([], []);
102  val extend = I;
103  fun merge
104    (Rules {rules = rules1, wrappers = (ws1, ws1'), ...},
105      Rules {rules = rules2, wrappers = (ws2, ws2'), ...}) =
106    let
107      val wrappers =
108        (Library.merge (eq_snd (op =)) (ws1, ws2), Library.merge (eq_snd (op =)) (ws1', ws2'));
109      val rules = Library.merge (fn ((_, (k1, th1)), (_, (k2, th2))) =>
110          k1 = k2 andalso Thm.eq_thm_prop (th1, th2)) (rules1, rules2);
111      val next = ~ (length rules);
112      val netpairs = fold (fn (n, (w, ((i, b), th))) =>
113          nth_map i (Tactic.insert_tagged_brl ((w, n), (b, th))))
114        (next upto ~1 ~~ rules) empty_netpairs;
115    in make_rules (next - 1) rules netpairs wrappers end;
116);
117
118fun print_rules ctxt =
119  let
120    val Rules {rules, ...} = Rules.get (Context.Proof ctxt);
121    fun prt_kind (i, b) =
122      Pretty.big_list ((the o AList.lookup (op =) kind_names) (i, b) ^ ":")
123        (map_filter (fn (_, (k, th)) =>
124            if k = (i, b) then SOME (Thm.pretty_thm_item ctxt th) else NONE)
125          (sort (int_ord o apply2 fst) rules));
126  in Pretty.writeln_chunks (map prt_kind rule_kinds) end;
127
128
129(* access data *)
130
131fun netpairs ctxt = let val Rules {netpairs, ...} = Rules.get (Context.Proof ctxt) in netpairs end;
132val netpair_bang = hd o netpairs;
133val netpair = hd o tl o netpairs;
134
135
136(* retrieving rules *)
137
138fun untaglist [] = []
139  | untaglist [(_ : int * int, x)] = [x]
140  | untaglist ((k, x) :: (rest as (k', _) :: _)) =
141      if k = k' then untaglist rest
142      else x :: untaglist rest;
143
144fun orderlist brls =
145  untaglist (sort (prod_ord int_ord int_ord o apply2 fst) brls);
146
147fun orderlist_no_weight brls =
148  untaglist (sort (int_ord o apply2 (snd o fst)) brls);
149
150local
151
152fun may_unify weighted t net =
153  map snd ((if weighted then orderlist else orderlist_no_weight) (Net.unify_term net t));
154
155fun find_erules _ [] = K []
156  | find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact));
157
158fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal);
159
160in
161
162fun find_rules_netpair ctxt weighted facts goal (inet, enet) =
163  find_erules weighted facts enet @ find_irules weighted goal inet
164  |> map (Thm.transfer' ctxt);
165
166fun find_rules ctxt weighted facts goal =
167  map (find_rules_netpair ctxt weighted facts goal) (netpairs ctxt);
168
169end;
170
171
172(* wrappers *)
173
174fun gen_add_wrapper upd w =
175  Context.theory_map (Rules.map (fn Rules {next, rules, netpairs, wrappers} =>
176    make_rules next rules netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers)));
177
178val addSWrapper = gen_add_wrapper Library.apfst;
179val addWrapper = gen_add_wrapper Library.apsnd;
180
181
182fun gen_wrap which ctxt =
183  let val Rules {wrappers, ...} = Rules.get (Context.Proof ctxt)
184  in fold_rev (fn (w, _) => w ctxt) (which wrappers) end;
185
186val Swrap = gen_wrap #1;
187val wrap = gen_wrap #2;
188
189
190
191(** attributes **)
192
193(* add and del rules *)
194
195
196val rule_del = Thm.declaration_attribute (Rules.map o del_rule);
197
198fun rule_add k view opt_w =
199  Thm.declaration_attribute (fn th => Rules.map (add_rule k opt_w (view th) o del_rule th));
200
201val intro_bang  = rule_add intro_bangK I;
202val elim_bang   = rule_add elim_bangK I;
203val dest_bang   = rule_add elim_bangK Tactic.make_elim;
204val intro       = rule_add introK I;
205val elim        = rule_add elimK I;
206val dest        = rule_add elimK Tactic.make_elim;
207val intro_query = rule_add intro_queryK I;
208val elim_query  = rule_add elim_queryK I;
209val dest_query  = rule_add elim_queryK Tactic.make_elim;
210
211val _ = Theory.setup
212  (snd o Global_Theory.add_thms [((Binding.empty, Drule.equal_intr_rule), [intro_query NONE])]);
213
214
215(* concrete syntax *)
216
217fun add a b c x =
218  (Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) --
219    Scan.option Parse.nat) >> (fn (f, n) => f n)) x;
220
221val _ = Theory.setup
222 (Attrib.setup \<^binding>\<open>intro\<close> (add intro_bang intro intro_query)
223    "declaration of introduction rule" #>
224  Attrib.setup \<^binding>\<open>elim\<close> (add elim_bang elim elim_query)
225    "declaration of elimination rule" #>
226  Attrib.setup \<^binding>\<open>dest\<close> (add dest_bang dest dest_query)
227    "declaration of destruction rule" #>
228  Attrib.setup \<^binding>\<open>rule\<close> (Scan.lift Args.del >> K rule_del)
229    "remove declaration of intro/elim/dest rule");
230
231end;
232