author | blanchet |
Fri, 01 Aug 2014 14:43:57 +0200 | |
changeset 57743 | 0af2d5dfb0ac |
parent 56334 | 6b3739fee456 |
child 59058 | a78612c67ec0 |
permissions | -rw-r--r-- |
12350 | 1 |
(* Title: Pure/Isar/context_rules.ML |
2 |
Author: Stefan Berghofer and Markus Wenzel, TU Muenchen |
|
3 |
||
18728 | 4 |
Declarations of intro/elim/dest rules in Pure (see also |
5 |
Provers/classical.ML for a more specialized version of the same idea). |
|
12350 | 6 |
*) |
7 |
||
8 |
signature CONTEXT_RULES = |
|
9 |
sig |
|
33370 | 10 |
type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net |
20289 | 11 |
val netpair_bang: Proof.context -> netpair |
12 |
val netpair: Proof.context -> netpair |
|
12350 | 13 |
val orderlist: ((int * int) * 'a) list -> 'a list |
12399 | 14 |
val find_rules_netpair: bool -> thm list -> term -> netpair -> thm list |
20289 | 15 |
val find_rules: bool -> thm list -> term -> Proof.context -> thm list list |
21506 | 16 |
val print_rules: Proof.context -> unit |
51798 | 17 |
val addSWrapper: (Proof.context -> (int -> tactic) -> int -> tactic) -> theory -> theory |
18 |
val addWrapper: (Proof.context -> (int -> tactic) -> int -> tactic) -> theory -> theory |
|
20289 | 19 |
val Swrap: Proof.context -> (int -> tactic) -> int -> tactic |
20 |
val wrap: Proof.context -> (int -> tactic) -> int -> tactic |
|
18728 | 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 |
|
30529 | 31 |
val add: (int option -> attribute) -> (int option -> attribute) -> (int option -> attribute) -> |
32 |
attribute context_parser |
|
12350 | 33 |
end; |
34 |
||
33369 | 35 |
structure Context_Rules: CONTEXT_RULES = |
12350 | 36 |
struct |
37 |
||
38 |
||
39 |
(** rule declaration contexts **) |
|
40 |
||
41 |
(* rule kinds *) |
|
42 |
||
43 |
val intro_bangK = (0, false); |
|
44 |
val elim_bangK = (0, true); |
|
45 |
val introK = (1, false); |
|
46 |
val elimK = (1, true); |
|
47 |
val intro_queryK = (2, false); |
|
48 |
val elim_queryK = (2, true); |
|
49 |
||
50 |
val 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 |
||
58 |
val rule_kinds = map #1 kind_names; |
|
19046
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
wenzelm
parents:
18977
diff
changeset
|
59 |
val rule_indexes = distinct (op =) (map #1 rule_kinds); |
12350 | 60 |
|
61 |
||
18637 | 62 |
(* context data *) |
12350 | 63 |
|
64 |
type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net; |
|
65 |
val empty_netpairs: netpair list = replicate (length rule_indexes) (Net.empty, Net.empty); |
|
66 |
||
33370 | 67 |
datatype rules = Rules of |
12350 | 68 |
{next: int, |
69 |
rules: (int * ((int * bool) * thm)) list, |
|
70 |
netpairs: netpair list, |
|
51798 | 71 |
wrappers: |
72 |
((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list * |
|
73 |
((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list}; |
|
12350 | 74 |
|
75 |
fun make_rules next rules netpairs wrappers = |
|
76 |
Rules {next = next, rules = rules, netpairs = netpairs, wrappers = wrappers}; |
|
77 |
||
78 |
fun add_rule (i, b) opt_w th (Rules {next, rules, netpairs, wrappers}) = |
|
15531 | 79 |
let val w = (case opt_w of SOME w => w | NONE => Tactic.subgoals_of_brl (b, th)) in |
12350 | 80 |
make_rules (next - 1) ((w, ((i, b), th)) :: rules) |
23227 | 81 |
(nth_map i (Tactic.insert_tagged_brl ((w, next), (b, th))) netpairs) wrappers |
12350 | 82 |
end; |
83 |
||
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
39557
diff
changeset
|
84 |
fun del_rule0 th (rs as Rules {next, rules, netpairs, wrappers}) = |
12350 | 85 |
let |
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
21506
diff
changeset
|
86 |
fun eq_th (_, (_, th')) = Thm.eq_thm_prop (th, th'); |
23227 | 87 |
fun del b netpair = Tactic.delete_tagged_brl (b, th) netpair handle Net.DELETE => netpair; |
12350 | 88 |
in |
89 |
if not (exists eq_th rules) then rs |
|
90 |
else make_rules next (filter_out eq_th rules) (map (del false o del true) netpairs) wrappers |
|
91 |
end; |
|
92 |
||
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
39557
diff
changeset
|
93 |
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
|
94 |
|
33519 | 95 |
structure Rules = Generic_Data |
18637 | 96 |
( |
33370 | 97 |
type T = rules; |
12350 | 98 |
val empty = make_rules ~1 [] empty_netpairs ([], []); |
16424 | 99 |
val extend = I; |
33519 | 100 |
fun merge |
101 |
(Rules {rules = rules1, wrappers = (ws1, ws1'), ...}, |
|
12350 | 102 |
Rules {rules = rules2, wrappers = (ws2, ws2'), ...}) = |
103 |
let |
|
18637 | 104 |
val wrappers = |
18921 | 105 |
(Library.merge (eq_snd (op =)) (ws1, ws2), Library.merge (eq_snd (op =)) (ws1', ws2')); |
106 |
val rules = Library.merge (fn ((_, (k1, th1)), (_, (k2, th2))) => |
|
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
21506
diff
changeset
|
107 |
k1 = k2 andalso Thm.eq_thm_prop (th1, th2)) (rules1, rules2); |
12350 | 108 |
val next = ~ (length rules); |
19473 | 109 |
val netpairs = fold (fn (n, (w, ((i, b), th))) => |
23227 | 110 |
nth_map i (Tactic.insert_tagged_brl ((w, n), (b, th)))) |
19473 | 111 |
(next upto ~1 ~~ rules) empty_netpairs; |
23227 | 112 |
in make_rules (next - 1) rules netpairs wrappers end; |
18637 | 113 |
); |
12350 | 114 |
|
22846 | 115 |
fun print_rules ctxt = |
116 |
let |
|
117 |
val Rules {rules, ...} = Rules.get (Context.Proof ctxt); |
|
118 |
fun prt_kind (i, b) = |
|
119 |
Pretty.big_list ((the o AList.lookup (op =) kind_names) (i, b) ^ ":") |
|
120 |
(map_filter (fn (_, (k, th)) => |
|
51584 | 121 |
if k = (i, b) then SOME (Display.pretty_thm_item ctxt th) else NONE) |
22846 | 122 |
(sort (int_ord o pairself fst) rules)); |
56334
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents:
56204
diff
changeset
|
123 |
in Pretty.writeln_chunks (map prt_kind rule_kinds) end; |
12350 | 124 |
|
125 |
||
126 |
(* access data *) |
|
127 |
||
18637 | 128 |
fun netpairs ctxt = let val Rules {netpairs, ...} = Rules.get (Context.Proof ctxt) in netpairs end; |
12350 | 129 |
val netpair_bang = hd o netpairs; |
130 |
val netpair = hd o tl o netpairs; |
|
131 |
||
132 |
||
12399 | 133 |
(* retrieving rules *) |
134 |
||
12350 | 135 |
fun untaglist [] = [] |
32784 | 136 |
| untaglist [(_ : int * int, x)] = [x] |
137 |
| untaglist ((k, x) :: (rest as (k', _) :: _)) = |
|
12350 | 138 |
if k = k' then untaglist rest |
139 |
else x :: untaglist rest; |
|
140 |
||
16424 | 141 |
fun orderlist brls = |
16512 | 142 |
untaglist (sort (prod_ord int_ord int_ord o pairself fst) brls); |
16424 | 143 |
|
144 |
fun orderlist_no_weight brls = |
|
16512 | 145 |
untaglist (sort (int_ord o pairself (snd o fst)) brls); |
12399 | 146 |
|
147 |
fun may_unify weighted t net = |
|
148 |
map snd ((if weighted then orderlist else orderlist_no_weight) (Net.unify_term net t)); |
|
149 |
||
150 |
fun find_erules _ [] = K [] |
|
12805 | 151 |
| find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact)); |
16424 | 152 |
|
12399 | 153 |
fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal); |
12380 | 154 |
|
12399 | 155 |
fun find_rules_netpair weighted facts goal (inet, enet) = |
156 |
find_erules weighted facts enet @ find_irules weighted goal inet; |
|
12380 | 157 |
|
16424 | 158 |
fun find_rules weighted facts goals = |
159 |
map (find_rules_netpair weighted facts goals) o netpairs; |
|
12350 | 160 |
|
161 |
||
12399 | 162 |
(* wrappers *) |
163 |
||
18667 | 164 |
fun gen_add_wrapper upd w = |
32784 | 165 |
Context.theory_map (Rules.map (fn Rules {next, rules, netpairs, wrappers} => |
18667 | 166 |
make_rules next rules netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers))); |
12350 | 167 |
|
168 |
val addSWrapper = gen_add_wrapper Library.apfst; |
|
169 |
val addWrapper = gen_add_wrapper Library.apsnd; |
|
170 |
||
171 |
||
172 |
fun gen_wrap which ctxt = |
|
18637 | 173 |
let val Rules {wrappers, ...} = Rules.get (Context.Proof ctxt) |
51798 | 174 |
in fold_rev (fn (w, _) => w ctxt) (which wrappers) end; |
12350 | 175 |
|
176 |
val Swrap = gen_wrap #1; |
|
177 |
val wrap = gen_wrap #2; |
|
178 |
||
179 |
||
180 |
||
181 |
(** attributes **) |
|
182 |
||
183 |
(* add and del rules *) |
|
184 |
||
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
39557
diff
changeset
|
185 |
|
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
39557
diff
changeset
|
186 |
val rule_del = Thm.declaration_attribute (fn th => Rules.map (del_rule th)); |
12350 | 187 |
|
18637 | 188 |
fun rule_add k view opt_w = |
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
39557
diff
changeset
|
189 |
Thm.declaration_attribute (fn th => Rules.map (add_rule k opt_w (view th) o del_rule th)); |
12350 | 190 |
|
18637 | 191 |
val intro_bang = rule_add intro_bangK I; |
192 |
val elim_bang = rule_add elim_bangK I; |
|
193 |
val dest_bang = rule_add elim_bangK Tactic.make_elim; |
|
194 |
val intro = rule_add introK I; |
|
195 |
val elim = rule_add elimK I; |
|
196 |
val dest = rule_add elimK Tactic.make_elim; |
|
197 |
val intro_query = rule_add intro_queryK I; |
|
198 |
val elim_query = rule_add elim_queryK I; |
|
199 |
val dest_query = rule_add elim_queryK Tactic.make_elim; |
|
12350 | 200 |
|
53171 | 201 |
val _ = Theory.setup |
202 |
(snd o Global_Theory.add_thms [((Binding.empty, Drule.equal_intr_rule), [intro_query NONE])]); |
|
18637 | 203 |
|
12350 | 204 |
|
18637 | 205 |
(* concrete syntax *) |
206 |
||
30529 | 207 |
fun add a b c x = |
27809
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
26463
diff
changeset
|
208 |
(Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) -- |
36950 | 209 |
Scan.option Parse.nat) >> (fn (f, n) => f n)) x; |
18637 | 210 |
|
53171 | 211 |
val _ = Theory.setup |
56204 | 212 |
(Attrib.setup @{binding intro} (add intro_bang intro intro_query) |
30529 | 213 |
"declaration of introduction rule" #> |
56204 | 214 |
Attrib.setup @{binding elim} (add elim_bang elim elim_query) |
30529 | 215 |
"declaration of elimination rule" #> |
56204 | 216 |
Attrib.setup @{binding dest} (add dest_bang dest dest_query) |
30529 | 217 |
"declaration of destruction rule" #> |
56204 | 218 |
Attrib.setup @{binding rule} (Scan.lift Args.del >> K rule_del) |
53171 | 219 |
"remove declaration of intro/elim/dest rule"); |
12350 | 220 |
|
221 |
end; |