| author | wenzelm | 
| Tue, 22 Oct 2024 19:26:40 +0200 | |
| changeset 81232 | 9867b5cf3f7a | 
| parent 77908 | a6bd716a6124 | 
| child 82587 | 7415414bd9d8 | 
| 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  | 
| 61049 | 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  | 
|
| 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}) =
 | 
|
| 61049 | 79  | 
let  | 
| 
77908
 
a6bd716a6124
tuned: concise combinators instead of bulky case-expressions;
 
wenzelm 
parents: 
74561 
diff
changeset
 | 
80  | 
val w = opt_w |> \<^if_none>\<open>Tactic.subgoals_of_brl (b, th)\<close>;  | 
| 61049 | 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  | 
|
| 12350 | 85  | 
end;  | 
86  | 
||
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
87  | 
fun del_rule0 th (rs as Rules {next, rules, netpairs, wrappers}) =
 | 
| 12350 | 88  | 
let  | 
| 
22360
 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 
wenzelm 
parents: 
21506 
diff
changeset
 | 
89  | 
fun eq_th (_, (_, th')) = Thm.eq_thm_prop (th, th');  | 
| 23227 | 90  | 
fun del b netpair = Tactic.delete_tagged_brl (b, th) netpair handle Net.DELETE => netpair;  | 
| 12350 | 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  | 
||
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
96  | 
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
 | 
97  | 
|
| 33519 | 98  | 
structure Rules = Generic_Data  | 
| 18637 | 99  | 
(  | 
| 33370 | 100  | 
type T = rules;  | 
| 12350 | 101  | 
val empty = make_rules ~1 [] empty_netpairs ([], []);  | 
| 33519 | 102  | 
fun merge  | 
103  | 
    (Rules {rules = rules1, wrappers = (ws1, ws1'), ...},
 | 
|
| 12350 | 104  | 
      Rules {rules = rules2, wrappers = (ws2, ws2'), ...}) =
 | 
105  | 
let  | 
|
| 18637 | 106  | 
val wrappers =  | 
| 18921 | 107  | 
(Library.merge (eq_snd (op =)) (ws1, ws2), Library.merge (eq_snd (op =)) (ws1', ws2'));  | 
108  | 
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
 | 
109  | 
k1 = k2 andalso Thm.eq_thm_prop (th1, th2)) (rules1, rules2);  | 
| 12350 | 110  | 
val next = ~ (length rules);  | 
| 19473 | 111  | 
val netpairs = fold (fn (n, (w, ((i, b), th))) =>  | 
| 23227 | 112  | 
nth_map i (Tactic.insert_tagged_brl ((w, n), (b, th))))  | 
| 19473 | 113  | 
(next upto ~1 ~~ rules) empty_netpairs;  | 
| 23227 | 114  | 
in make_rules (next - 1) rules netpairs wrappers end;  | 
| 18637 | 115  | 
);  | 
| 12350 | 116  | 
|
| 22846 | 117  | 
fun print_rules ctxt =  | 
118  | 
let  | 
|
119  | 
    val Rules {rules, ...} = Rules.get (Context.Proof ctxt);
 | 
|
120  | 
fun prt_kind (i, b) =  | 
|
121  | 
Pretty.big_list ((the o AList.lookup (op =) kind_names) (i, b) ^ ":")  | 
|
122  | 
(map_filter (fn (_, (k, th)) =>  | 
|
| 61268 | 123  | 
if k = (i, b) then SOME (Thm.pretty_thm_item ctxt th) else NONE)  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
56334 
diff
changeset
 | 
124  | 
(sort (int_ord o apply2 fst) rules));  | 
| 
56334
 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 
wenzelm 
parents: 
56204 
diff
changeset
 | 
125  | 
in Pretty.writeln_chunks (map prt_kind rule_kinds) end;  | 
| 12350 | 126  | 
|
127  | 
||
128  | 
(* access data *)  | 
|
129  | 
||
| 18637 | 130  | 
fun netpairs ctxt = let val Rules {netpairs, ...} = Rules.get (Context.Proof ctxt) in netpairs end;
 | 
| 12350 | 131  | 
val netpair_bang = hd o netpairs;  | 
132  | 
val netpair = hd o tl o netpairs;  | 
|
133  | 
||
134  | 
||
| 12399 | 135  | 
(* retrieving rules *)  | 
136  | 
||
| 12350 | 137  | 
fun untaglist [] = []  | 
| 32784 | 138  | 
| untaglist [(_ : int * int, x)] = [x]  | 
139  | 
| untaglist ((k, x) :: (rest as (k', _) :: _)) =  | 
|
| 12350 | 140  | 
if k = k' then untaglist rest  | 
141  | 
else x :: untaglist rest;  | 
|
142  | 
||
| 16424 | 143  | 
fun orderlist brls =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
56334 
diff
changeset
 | 
144  | 
untaglist (sort (prod_ord int_ord int_ord o apply2 fst) brls);  | 
| 16424 | 145  | 
|
146  | 
fun orderlist_no_weight brls =  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
56334 
diff
changeset
 | 
147  | 
untaglist (sort (int_ord o apply2 (snd o fst)) brls);  | 
| 12399 | 148  | 
|
| 61049 | 149  | 
local  | 
150  | 
||
| 12399 | 151  | 
fun may_unify weighted t net =  | 
152  | 
map snd ((if weighted then orderlist else orderlist_no_weight) (Net.unify_term net t));  | 
|
153  | 
||
154  | 
fun find_erules _ [] = K []  | 
|
| 12805 | 155  | 
| find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact));  | 
| 16424 | 156  | 
|
| 12399 | 157  | 
fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal);  | 
| 12380 | 158  | 
|
| 61049 | 159  | 
in  | 
| 12380 | 160  | 
|
| 61049 | 161  | 
fun find_rules_netpair ctxt weighted facts goal (inet, enet) =  | 
162  | 
find_erules weighted facts enet @ find_irules weighted goal inet  | 
|
| 67649 | 163  | 
|> map (Thm.transfer' ctxt);  | 
| 61049 | 164  | 
|
165  | 
fun find_rules ctxt weighted facts goal =  | 
|
166  | 
map (find_rules_netpair ctxt weighted facts goal) (netpairs ctxt);  | 
|
167  | 
||
168  | 
end;  | 
|
| 12350 | 169  | 
|
170  | 
||
| 12399 | 171  | 
(* wrappers *)  | 
172  | 
||
| 18667 | 173  | 
fun gen_add_wrapper upd w =  | 
| 32784 | 174  | 
  Context.theory_map (Rules.map (fn Rules {next, rules, netpairs, wrappers} =>
 | 
| 18667 | 175  | 
make_rules next rules netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers)));  | 
| 12350 | 176  | 
|
177  | 
val addSWrapper = gen_add_wrapper Library.apfst;  | 
|
178  | 
val addWrapper = gen_add_wrapper Library.apsnd;  | 
|
179  | 
||
180  | 
||
181  | 
fun gen_wrap which ctxt =  | 
|
| 18637 | 182  | 
  let val Rules {wrappers, ...} = Rules.get (Context.Proof ctxt)
 | 
| 51798 | 183  | 
in fold_rev (fn (w, _) => w ctxt) (which wrappers) end;  | 
| 12350 | 184  | 
|
185  | 
val Swrap = gen_wrap #1;  | 
|
186  | 
val wrap = gen_wrap #2;  | 
|
187  | 
||
188  | 
||
189  | 
||
190  | 
(** attributes **)  | 
|
191  | 
||
192  | 
(* add and del rules *)  | 
|
193  | 
||
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
39557 
diff
changeset
 | 
194  | 
|
| 67626 | 195  | 
val rule_del = Thm.declaration_attribute (Rules.map o del_rule);  | 
| 12350 | 196  | 
|
| 18637 | 197  | 
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
 | 
198  | 
Thm.declaration_attribute (fn th => Rules.map (add_rule k opt_w (view th) o del_rule th));  | 
| 12350 | 199  | 
|
| 18637 | 200  | 
val intro_bang = rule_add intro_bangK I;  | 
201  | 
val elim_bang = rule_add elim_bangK I;  | 
|
202  | 
val dest_bang = rule_add elim_bangK Tactic.make_elim;  | 
|
203  | 
val intro = rule_add introK I;  | 
|
204  | 
val elim = rule_add elimK I;  | 
|
205  | 
val dest = rule_add elimK Tactic.make_elim;  | 
|
206  | 
val intro_query = rule_add intro_queryK I;  | 
|
207  | 
val elim_query = rule_add elim_queryK I;  | 
|
208  | 
val dest_query = rule_add elim_queryK Tactic.make_elim;  | 
|
| 12350 | 209  | 
|
| 53171 | 210  | 
val _ = Theory.setup  | 
211  | 
(snd o Global_Theory.add_thms [((Binding.empty, Drule.equal_intr_rule), [intro_query NONE])]);  | 
|
| 18637 | 212  | 
|
| 12350 | 213  | 
|
| 18637 | 214  | 
(* concrete syntax *)  | 
215  | 
||
| 30529 | 216  | 
fun add a b c x =  | 
| 
27809
 
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
217  | 
(Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) --  | 
| 36950 | 218  | 
Scan.option Parse.nat) >> (fn (f, n) => f n)) x;  | 
| 18637 | 219  | 
|
| 53171 | 220  | 
val _ = Theory.setup  | 
| 67147 | 221  | 
(Attrib.setup \<^binding>\<open>intro\<close> (add intro_bang intro intro_query)  | 
| 30529 | 222  | 
"declaration of introduction rule" #>  | 
| 67147 | 223  | 
Attrib.setup \<^binding>\<open>elim\<close> (add elim_bang elim elim_query)  | 
| 30529 | 224  | 
"declaration of elimination rule" #>  | 
| 67147 | 225  | 
Attrib.setup \<^binding>\<open>dest\<close> (add dest_bang dest dest_query)  | 
| 30529 | 226  | 
"declaration of destruction rule" #>  | 
| 67147 | 227  | 
Attrib.setup \<^binding>\<open>rule\<close> (Scan.lift Args.del >> K rule_del)  | 
| 53171 | 228  | 
"remove declaration of intro/elim/dest rule");  | 
| 12350 | 229  | 
|
230  | 
end;  |