author | haftmann |
Sun, 27 Jul 2025 17:52:06 +0200 | |
changeset 82909 | e4fae2227594 |
parent 82874 | abfb6ed8ec21 |
permissions | -rw-r--r-- |
12350 | 1 |
(* Title: Pure/Isar/context_rules.ML |
82816 | 2 |
Author: Stefan Berghofer, TU Muenchen |
3 |
Author: Makarius |
|
12350 | 4 |
|
18728 | 5 |
Declarations of intro/elim/dest rules in Pure (see also |
6 |
Provers/classical.ML for a more specialized version of the same idea). |
|
12350 | 7 |
*) |
8 |
||
9 |
signature CONTEXT_RULES = |
|
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 | 14 |
val find_rules: Proof.context -> bool -> thm list -> term -> thm list list |
82874 | 15 |
val declared_rule: Context.generic -> thm -> bool |
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 |
||
82839
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
41 |
(* rule declarations *) |
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
42 |
|
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
43 |
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
|
44 |
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
|
45 |
|
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
46 |
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
|
47 |
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
|
48 |
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
|
49 |
|
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
50 |
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
|
51 |
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
|
52 |
|
82848
0a8977dcb21a
more accurate delete operation via authentic index --- minor performance tuning;
wenzelm
parents:
82847
diff
changeset
|
53 |
fun remove_decl ({tag = {index, ...}, info = th, ...}: decl) = |
0a8977dcb21a
more accurate delete operation via authentic index --- minor performance tuning;
wenzelm
parents:
82847
diff
changeset
|
54 |
Bires.delete_tagged_rule (index, th); |
82839
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
55 |
|
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
56 |
|
18637 | 57 |
(* context data *) |
12350 | 58 |
|
33370 | 59 |
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
|
60 |
{decls: decls, |
82812
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents:
82808
diff
changeset
|
61 |
netpairs: Bires.netpair list, |
51798 | 62 |
wrappers: |
63 |
((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list * |
|
64 |
((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list}; |
|
12350 | 65 |
|
82826
f5fd9b41188a
efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents:
82818
diff
changeset
|
66 |
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
|
67 |
Rules {decls = decls, netpairs = netpairs, wrappers = wrappers}; |
12350 | 68 |
|
82826
f5fd9b41188a
efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents:
82818
diff
changeset
|
69 |
fun add_rule kind opt_weight th (rules as Rules {decls, netpairs, wrappers}) = |
61049 | 70 |
let |
82839
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
71 |
val th' = Thm.trim_context th; |
82873 | 72 |
val th'' = Thm.trim_context (Bires.kind_make_elim kind th); |
82868
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
73 |
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
|
74 |
in |
82839
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
75 |
(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
|
76 |
(NONE, _) => rules |
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
77 |
| (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
|
78 |
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
|
79 |
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
|
80 |
end; |
12350 | 81 |
|
82868
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
82 |
fun del_rule th (rules as Rules {decls, netpairs, wrappers}) = |
82826
f5fd9b41188a
efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents:
82818
diff
changeset
|
83 |
(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
|
84 |
([], _) => rules |
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
85 |
| (old_decls, decls') => |
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
86 |
let |
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
87 |
val netpairs' = |
60ec2b2dc55a
clarified signature: rule declarations work via "info" as internal rule (which coincides with external rule);
wenzelm
parents:
82836
diff
changeset
|
88 |
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
|
89 |
old_decls netpairs; |
82826
f5fd9b41188a
efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents:
82818
diff
changeset
|
90 |
in make_rules decls' netpairs' wrappers end); |
12350 | 91 |
|
82874 | 92 |
|
82816 | 93 |
structure Data = Generic_Data |
18637 | 94 |
( |
33370 | 95 |
type T = rules; |
82828 | 96 |
val empty = make_rules Bires.empty_decls (Bires.kind_values Bires.empty_netpair) ([], []); |
82847 | 97 |
fun merge (rules1, rules2) = |
98 |
if pointer_eq (rules1, rules2) then rules1 |
|
99 |
else |
|
12350 | 100 |
let |
82847 | 101 |
val Rules {decls = decls1, netpairs = netpairs1, wrappers = (ws1, ws1')} = rules1; |
102 |
val Rules {decls = decls2, netpairs = _, wrappers = (ws2, ws2')} = rules2; |
|
82826
f5fd9b41188a
efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents:
82818
diff
changeset
|
103 |
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
|
104 |
val netpairs = |
f5fd9b41188a
efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents:
82818
diff
changeset
|
105 |
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
|
106 |
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
|
107 |
if Bires.kind_index (#kind decl) = i then insert_decl decl else I))) |
18637 | 108 |
val wrappers = |
82826
f5fd9b41188a
efficient rule declarations in canonical order, for update of netpairs and print operation;
wenzelm
parents:
82818
diff
changeset
|
109 |
(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
|
110 |
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
|
111 |
in make_rules decls netpairs wrappers end; |
18637 | 112 |
); |
12350 | 113 |
|
82874 | 114 |
fun declared_rule context = |
115 |
let val Rules {decls, ...} = Data.get context |
|
116 |
in Bires.has_decls decls end; |
|
117 |
||
22846 | 118 |
fun print_rules ctxt = |
82836 | 119 |
let val Rules {decls, ...} = Data.get (Context.Proof ctxt) |
82864 | 120 |
in Pretty.writeln (Pretty.chunks (Bires.pretty_decls ctxt decls)) end; |
12350 | 121 |
|
122 |
||
123 |
(* access data *) |
|
124 |
||
82816 | 125 |
fun netpairs ctxt = let val Rules {netpairs, ...} = Data.get (Context.Proof ctxt) in netpairs end; |
12350 | 126 |
val netpair_bang = hd o netpairs; |
127 |
val netpair = hd o tl o netpairs; |
|
128 |
||
129 |
||
12399 | 130 |
(* retrieving rules *) |
131 |
||
61049 | 132 |
local |
133 |
||
82812
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents:
82808
diff
changeset
|
134 |
fun order weighted = |
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents:
82808
diff
changeset
|
135 |
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
|
136 |
|
12399 | 137 |
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
|
138 |
map snd (order weighted (Net.unify_term net t)); |
12399 | 139 |
|
140 |
fun find_erules _ [] = K [] |
|
12805 | 141 |
| find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact)); |
16424 | 142 |
|
12399 | 143 |
fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal); |
12380 | 144 |
|
61049 | 145 |
in |
12380 | 146 |
|
61049 | 147 |
fun find_rules_netpair ctxt weighted facts goal (inet, enet) = |
148 |
find_erules weighted facts enet @ find_irules weighted goal inet |
|
67649 | 149 |
|> map (Thm.transfer' ctxt); |
61049 | 150 |
|
151 |
fun find_rules ctxt weighted facts goal = |
|
152 |
map (find_rules_netpair ctxt weighted facts goal) (netpairs ctxt); |
|
153 |
||
154 |
end; |
|
12350 | 155 |
|
156 |
||
12399 | 157 |
(* wrappers *) |
158 |
||
18667 | 159 |
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
|
160 |
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
|
161 |
make_rules decls netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers))); |
12350 | 162 |
|
163 |
val addSWrapper = gen_add_wrapper Library.apfst; |
|
164 |
val addWrapper = gen_add_wrapper Library.apsnd; |
|
165 |
||
166 |
||
167 |
fun gen_wrap which ctxt = |
|
82816 | 168 |
let val Rules {wrappers, ...} = Data.get (Context.Proof ctxt) |
51798 | 169 |
in fold_rev (fn (w, _) => w ctxt) (which wrappers) end; |
12350 | 170 |
|
171 |
val Swrap = gen_wrap #1; |
|
172 |
val wrap = gen_wrap #2; |
|
173 |
||
174 |
||
175 |
||
176 |
(** attributes **) |
|
177 |
||
178 |
(* add and del rules *) |
|
179 |
||
82816 | 180 |
val rule_del = Thm.declaration_attribute (Data.map o del_rule); |
12350 | 181 |
|
82868
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
182 |
fun rule_add k opt_w = |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
183 |
Thm.declaration_attribute (fn th => Data.map (add_rule k opt_w th o del_rule th)); |
12350 | 184 |
|
82868
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
185 |
val intro_bang = rule_add Bires.intro_bang_kind; |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
186 |
val elim_bang = rule_add Bires.elim_bang_kind; |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
187 |
val dest_bang = rule_add Bires.dest_bang_kind; |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
188 |
val intro = rule_add Bires.intro_kind; |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
189 |
val elim = rule_add Bires.elim_kind; |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
190 |
val dest = rule_add Bires.dest_kind; |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
191 |
val intro_query = rule_add Bires.intro_query_kind; |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
192 |
val elim_query = rule_add Bires.elim_query_kind; |
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
wenzelm
parents:
82864
diff
changeset
|
193 |
val dest_query = rule_add Bires.dest_query_kind; |
12350 | 194 |
|
53171 | 195 |
val _ = Theory.setup |
196 |
(snd o Global_Theory.add_thms [((Binding.empty, Drule.equal_intr_rule), [intro_query NONE])]); |
|
18637 | 197 |
|
12350 | 198 |
|
18637 | 199 |
(* concrete syntax *) |
200 |
||
30529 | 201 |
fun add a b c x = |
27809
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
26463
diff
changeset
|
202 |
(Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) -- |
36950 | 203 |
Scan.option Parse.nat) >> (fn (f, n) => f n)) x; |
18637 | 204 |
|
53171 | 205 |
val _ = Theory.setup |
67147 | 206 |
(Attrib.setup \<^binding>\<open>intro\<close> (add intro_bang intro intro_query) |
82818 | 207 |
"declaration of Pure introduction rule" #> |
67147 | 208 |
Attrib.setup \<^binding>\<open>elim\<close> (add elim_bang elim elim_query) |
82818 | 209 |
"declaration of Pure elimination rule" #> |
67147 | 210 |
Attrib.setup \<^binding>\<open>dest\<close> (add dest_bang dest dest_query) |
82818 | 211 |
"declaration of Pure destruction rule" #> |
67147 | 212 |
Attrib.setup \<^binding>\<open>rule\<close> (Scan.lift Args.del >> K rule_del) |
82818 | 213 |
"remove declaration of Pure intro/elim/dest rule"); |
12350 | 214 |
|
215 |
end; |