12350
|
1 |
(* Title: Pure/Isar/context_rules.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Stefan Berghofer and Markus Wenzel, TU Muenchen
|
|
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE)
|
|
5 |
|
|
6 |
Declarations of intro/elim/dest rules in Pure; see
|
|
7 |
Provers/classical.ML for a more specialized version of the same idea.
|
|
8 |
*)
|
|
9 |
|
|
10 |
signature CONTEXT_RULES =
|
|
11 |
sig
|
|
12 |
type netpair
|
|
13 |
type T
|
|
14 |
val netpair_bang: Proof.context -> netpair
|
|
15 |
val netpair: Proof.context -> netpair
|
|
16 |
val orderlist: ((int * int) * 'a) list -> 'a list
|
12399
|
17 |
val find_rules_netpair: bool -> thm list -> term -> netpair -> thm list
|
|
18 |
val find_rules: bool -> thm list -> term -> Proof.context -> thm list list
|
12350
|
19 |
val print_global_rules: theory -> unit
|
|
20 |
val print_local_rules: Proof.context -> unit
|
|
21 |
val addSWrapper: ((int -> tactic) -> int -> tactic) -> theory -> theory
|
|
22 |
val addWrapper: ((int -> tactic) -> int -> tactic) -> theory -> theory
|
|
23 |
val Swrap: Proof.context -> (int -> tactic) -> int -> tactic
|
|
24 |
val wrap: Proof.context -> (int -> tactic) -> int -> tactic
|
|
25 |
val intro_bang_global: int option -> theory attribute
|
|
26 |
val elim_bang_global: int option -> theory attribute
|
|
27 |
val dest_bang_global: int option -> theory attribute
|
|
28 |
val intro_global: int option -> theory attribute
|
|
29 |
val elim_global: int option -> theory attribute
|
|
30 |
val dest_global: int option -> theory attribute
|
|
31 |
val intro_query_global: int option -> theory attribute
|
|
32 |
val elim_query_global: int option -> theory attribute
|
|
33 |
val dest_query_global: int option -> theory attribute
|
|
34 |
val rule_del_global: theory attribute
|
|
35 |
val intro_bang_local: int option -> Proof.context attribute
|
|
36 |
val elim_bang_local: int option -> Proof.context attribute
|
|
37 |
val dest_bang_local: int option -> Proof.context attribute
|
|
38 |
val intro_local: int option -> Proof.context attribute
|
|
39 |
val elim_local: int option -> Proof.context attribute
|
|
40 |
val dest_local: int option -> Proof.context attribute
|
|
41 |
val intro_query_local: int option -> Proof.context attribute
|
|
42 |
val elim_query_local: int option -> Proof.context attribute
|
|
43 |
val dest_query_local: int option -> Proof.context attribute
|
|
44 |
val rule_del_local: Proof.context attribute
|
12380
|
45 |
val addXIs_global: theory * thm list -> theory
|
|
46 |
val addXEs_global: theory * thm list -> theory
|
|
47 |
val addXDs_global: theory * thm list -> theory
|
|
48 |
val delrules_global: theory * thm list -> theory
|
|
49 |
val addXIs_local: Proof.context * thm list -> Proof.context
|
|
50 |
val addXEs_local: Proof.context * thm list -> Proof.context
|
|
51 |
val addXDs_local: Proof.context * thm list -> Proof.context
|
|
52 |
val delrules_local: Proof.context * thm list -> Proof.context
|
12350
|
53 |
val setup: (theory -> theory) list
|
|
54 |
end;
|
|
55 |
|
|
56 |
structure ContextRules: CONTEXT_RULES =
|
|
57 |
struct
|
|
58 |
|
|
59 |
|
|
60 |
(** rule declaration contexts **)
|
|
61 |
|
|
62 |
(* rule kinds *)
|
|
63 |
|
|
64 |
val intro_bangK = (0, false);
|
|
65 |
val elim_bangK = (0, true);
|
|
66 |
val introK = (1, false);
|
|
67 |
val elimK = (1, true);
|
|
68 |
val intro_queryK = (2, false);
|
|
69 |
val elim_queryK = (2, true);
|
|
70 |
|
|
71 |
val kind_names =
|
|
72 |
[(intro_bangK, "safe introduction rules (intro!)"),
|
|
73 |
(elim_bangK, "safe elimination rules (elim!)"),
|
|
74 |
(introK, "introduction rules (intro)"),
|
|
75 |
(elimK, "elimination rules (elim)"),
|
|
76 |
(intro_queryK, "extra introduction rules (intro?)"),
|
|
77 |
(elim_queryK, "extra elimination rules (elim?)")];
|
|
78 |
|
|
79 |
val rule_kinds = map #1 kind_names;
|
|
80 |
val rule_indexes = distinct (map #1 rule_kinds);
|
|
81 |
|
|
82 |
|
|
83 |
(* raw data *)
|
|
84 |
|
|
85 |
type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net;
|
|
86 |
val empty_netpairs: netpair list = replicate (length rule_indexes) (Net.empty, Net.empty);
|
|
87 |
|
|
88 |
datatype T = Rules of
|
|
89 |
{next: int,
|
|
90 |
rules: (int * ((int * bool) * thm)) list,
|
|
91 |
netpairs: netpair list,
|
|
92 |
wrappers: (((int -> tactic) -> int -> tactic) * stamp) list *
|
|
93 |
(((int -> tactic) -> int -> tactic) * stamp) list};
|
|
94 |
|
|
95 |
fun make_rules next rules netpairs wrappers =
|
|
96 |
Rules {next = next, rules = rules, netpairs = netpairs, wrappers = wrappers};
|
|
97 |
|
|
98 |
fun add_rule (i, b) opt_w th (Rules {next, rules, netpairs, wrappers}) =
|
|
99 |
let val w = (case opt_w of Some w => w | None => Tactic.subgoals_of_brl (b, th)) in
|
|
100 |
make_rules (next - 1) ((w, ((i, b), th)) :: rules)
|
|
101 |
(map_nth_elem i (curry insert_tagged_brl ((w, next), (b, th))) netpairs) wrappers
|
|
102 |
end;
|
|
103 |
|
|
104 |
fun del_rule th (rs as Rules {next, rules, netpairs, wrappers}) =
|
|
105 |
let
|
|
106 |
fun eq_th (_, (_, th')) = Thm.eq_thm (th, th');
|
|
107 |
fun del b netpair = delete_tagged_brl ((b, th), netpair) handle Net.DELETE => netpair;
|
|
108 |
in
|
|
109 |
if not (exists eq_th rules) then rs
|
|
110 |
else make_rules next (filter_out eq_th rules) (map (del false o del true) netpairs) wrappers
|
|
111 |
end;
|
|
112 |
|
|
113 |
fun print_rules prt x (Rules {rules, ...}) =
|
|
114 |
let
|
|
115 |
fun prt_kind (i, b) =
|
|
116 |
Pretty.big_list (the (assoc (kind_names, (i, b))) ^ ":")
|
|
117 |
(mapfilter (fn (_, (k, th)) => if k = (i, b) then Some (prt x th) else None)
|
|
118 |
(sort (int_ord o pairself fst) rules));
|
|
119 |
in Pretty.writeln (Pretty.chunks (map prt_kind rule_kinds)) end;
|
|
120 |
|
|
121 |
|
|
122 |
(* theory and proof data *)
|
|
123 |
|
|
124 |
structure GlobalRulesArgs =
|
|
125 |
struct
|
|
126 |
val name = "Isar/rule_context";
|
|
127 |
type T = T;
|
|
128 |
|
|
129 |
val empty = make_rules ~1 [] empty_netpairs ([], []);
|
|
130 |
val copy = I;
|
|
131 |
val prep_ext = I;
|
|
132 |
|
|
133 |
fun merge (Rules {rules = rules1, wrappers = (ws1, ws1'), ...},
|
|
134 |
Rules {rules = rules2, wrappers = (ws2, ws2'), ...}) =
|
|
135 |
let
|
|
136 |
val wrappers = (gen_merge_lists' eq_snd ws1 ws2, gen_merge_lists' eq_snd ws1' ws2');
|
|
137 |
val rules = gen_merge_lists' (fn ((_, (k1, th1)), (_, (k2, th2))) =>
|
|
138 |
k1 = k2 andalso Thm.eq_thm (th1, th2)) rules1 rules2;
|
|
139 |
val next = ~ (length rules);
|
|
140 |
val netpairs = foldl (fn (nps, (n, (w, ((i, b), th)))) =>
|
|
141 |
map_nth_elem i (curry insert_tagged_brl ((w, n), (b, th))) nps)
|
|
142 |
(empty_netpairs, next upto ~1 ~~ rules);
|
|
143 |
in make_rules (next - 1) rules netpairs wrappers end;
|
|
144 |
|
|
145 |
val print = print_rules Display.pretty_thm_sg;
|
|
146 |
end;
|
|
147 |
|
|
148 |
structure GlobalRules = TheoryDataFun(GlobalRulesArgs);
|
|
149 |
val print_global_rules = GlobalRules.print;
|
|
150 |
|
|
151 |
structure LocalRulesArgs =
|
|
152 |
struct
|
|
153 |
val name = GlobalRulesArgs.name;
|
|
154 |
type T = GlobalRulesArgs.T;
|
|
155 |
val init = GlobalRules.get;
|
|
156 |
val print = print_rules ProofContext.pretty_thm;
|
|
157 |
end;
|
|
158 |
|
|
159 |
structure LocalRules = ProofDataFun(LocalRulesArgs);
|
|
160 |
val print_local_rules = LocalRules.print;
|
|
161 |
|
|
162 |
|
|
163 |
(* access data *)
|
|
164 |
|
|
165 |
fun netpairs ctxt = let val Rules {netpairs, ...} = LocalRules.get ctxt in netpairs end;
|
|
166 |
val netpair_bang = hd o netpairs;
|
|
167 |
val netpair = hd o tl o netpairs;
|
|
168 |
|
|
169 |
|
12399
|
170 |
(* retrieving rules *)
|
|
171 |
|
12350
|
172 |
fun untaglist [] = []
|
|
173 |
| untaglist [(k : int * int, x)] = [x]
|
|
174 |
| untaglist ((k, x) :: (rest as (k', x') :: _)) =
|
|
175 |
if k = k' then untaglist rest
|
|
176 |
else x :: untaglist rest;
|
|
177 |
|
|
178 |
fun orderlist brls = untaglist (sort (prod_ord int_ord int_ord o pairself fst) brls);
|
|
179 |
fun orderlist_no_weight brls = untaglist (sort (int_ord o pairself (snd o fst)) brls);
|
12399
|
180 |
|
|
181 |
fun may_unify weighted t net =
|
|
182 |
map snd ((if weighted then orderlist else orderlist_no_weight) (Net.unify_term net t));
|
|
183 |
|
|
184 |
fun find_erules _ [] = K []
|
12805
|
185 |
| find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact));
|
12399
|
186 |
fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal);
|
12380
|
187 |
|
12399
|
188 |
fun find_rules_netpair weighted facts goal (inet, enet) =
|
|
189 |
find_erules weighted facts enet @ find_irules weighted goal inet;
|
12380
|
190 |
|
12399
|
191 |
fun find_rules weighted facts goals = map (find_rules_netpair weighted facts goals) o netpairs;
|
12350
|
192 |
|
|
193 |
|
12399
|
194 |
(* wrappers *)
|
|
195 |
|
12350
|
196 |
fun gen_add_wrapper upd w = GlobalRules.map (fn (rs as Rules {next, rules, netpairs, wrappers}) =>
|
|
197 |
make_rules next rules netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers));
|
|
198 |
|
|
199 |
val addSWrapper = gen_add_wrapper Library.apfst;
|
|
200 |
val addWrapper = gen_add_wrapper Library.apsnd;
|
|
201 |
|
|
202 |
|
|
203 |
fun gen_wrap which ctxt =
|
|
204 |
let val Rules {wrappers, ...} = LocalRules.get ctxt
|
|
205 |
in fn tac => foldr (fn ((w, _), t) => w t) (which wrappers, tac) end;
|
|
206 |
|
|
207 |
val Swrap = gen_wrap #1;
|
|
208 |
val wrap = gen_wrap #2;
|
|
209 |
|
|
210 |
|
|
211 |
|
|
212 |
(** attributes **)
|
|
213 |
|
|
214 |
(* add and del rules *)
|
|
215 |
|
|
216 |
local
|
|
217 |
|
|
218 |
fun del map_data (x, th) =
|
|
219 |
(map_data (del_rule th o del_rule (Tactic.make_elim th)) x, th);
|
|
220 |
|
|
221 |
fun add k view map_data opt_w =
|
12412
|
222 |
(fn (x, th) => (map_data (add_rule k opt_w (view th)) x, th)) o del map_data;
|
12350
|
223 |
|
|
224 |
in
|
|
225 |
|
|
226 |
val intro_bang_global = add intro_bangK I GlobalRules.map;
|
|
227 |
val elim_bang_global = add elim_bangK I GlobalRules.map;
|
|
228 |
val dest_bang_global = add elim_bangK Tactic.make_elim GlobalRules.map;
|
|
229 |
val intro_global = add introK I GlobalRules.map;
|
|
230 |
val elim_global = add elimK I GlobalRules.map;
|
|
231 |
val dest_global = add elimK Tactic.make_elim GlobalRules.map;
|
|
232 |
val intro_query_global = add intro_queryK I GlobalRules.map;
|
|
233 |
val elim_query_global = add elim_queryK I GlobalRules.map;
|
|
234 |
val dest_query_global = add elim_queryK Tactic.make_elim GlobalRules.map;
|
|
235 |
val rule_del_global = del GlobalRules.map;
|
|
236 |
|
|
237 |
val intro_bang_local = add intro_bangK I LocalRules.map;
|
|
238 |
val elim_bang_local = add elim_bangK I LocalRules.map;
|
|
239 |
val dest_bang_local = add elim_bangK Tactic.make_elim LocalRules.map;
|
|
240 |
val intro_local = add introK I LocalRules.map;
|
|
241 |
val elim_local = add elimK I LocalRules.map;
|
|
242 |
val dest_local = add elimK Tactic.make_elim LocalRules.map;
|
|
243 |
val intro_query_local = add intro_queryK I LocalRules.map;
|
|
244 |
val elim_query_local = add elim_queryK I LocalRules.map;
|
|
245 |
val dest_query_local = add elim_queryK Tactic.make_elim LocalRules.map;
|
|
246 |
val rule_del_local = del LocalRules.map;
|
|
247 |
|
|
248 |
end;
|
|
249 |
|
|
250 |
|
12380
|
251 |
(* low-level modifiers *)
|
|
252 |
|
|
253 |
fun modifier att (x, ths) =
|
|
254 |
#1 (Thm.applys_attributes ((x, rev ths), [att]));
|
|
255 |
|
|
256 |
val addXIs_global = modifier (intro_query_global None);
|
|
257 |
val addXEs_global = modifier (elim_query_global None);
|
|
258 |
val addXDs_global = modifier (dest_query_global None);
|
|
259 |
val delrules_global = modifier rule_del_global;
|
|
260 |
|
|
261 |
val addXIs_local = modifier (intro_query_local None);
|
|
262 |
val addXEs_local = modifier (elim_query_local None);
|
|
263 |
val addXDs_local = modifier (dest_query_local None);
|
|
264 |
val delrules_local = modifier rule_del_local;
|
|
265 |
|
|
266 |
|
12350
|
267 |
(* concrete syntax *)
|
|
268 |
|
|
269 |
fun add_args a b c x = Attrib.syntax
|
12380
|
270 |
(Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) -- (Scan.option Args.nat))
|
|
271 |
>> (fn (f, n) => f n)) x;
|
12350
|
272 |
|
|
273 |
fun del_args att = Attrib.syntax (Scan.lift Args.del >> K att);
|
|
274 |
|
|
275 |
|
|
276 |
val rule_atts =
|
|
277 |
[("intro",
|
|
278 |
(add_args intro_bang_global intro_global intro_query_global,
|
|
279 |
add_args intro_bang_local intro_local intro_query_local),
|
|
280 |
"declaration of introduction rule"),
|
|
281 |
("elim",
|
|
282 |
(add_args elim_bang_global elim_global elim_query_global,
|
|
283 |
add_args elim_bang_local elim_local elim_query_local),
|
|
284 |
"declaration of elimination rule"),
|
|
285 |
("dest",
|
|
286 |
(add_args dest_bang_global dest_global dest_query_global,
|
|
287 |
add_args dest_bang_local dest_local dest_query_local),
|
|
288 |
"declaration of destruction rule"),
|
|
289 |
("rule", (del_args rule_del_global, del_args rule_del_local),
|
|
290 |
"remove declaration of intro/elim/dest rule")];
|
|
291 |
|
|
292 |
|
|
293 |
|
|
294 |
(** theory setup **)
|
|
295 |
|
|
296 |
val setup =
|
|
297 |
[GlobalRules.init, LocalRules.init, Attrib.add_attributes rule_atts];
|
|
298 |
|
|
299 |
|
|
300 |
end;
|