author | wenzelm |
Sun, 06 Jul 2025 14:58:00 +0200 | |
changeset 82818 | c6b3f0ee0a69 |
parent 82817 | be1fb22d9e2a |
child 82826 | f5fd9b41188a |
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 |
21506 | 15 |
val print_rules: Proof.context -> unit |
51798 | 16 |
val addSWrapper: (Proof.context -> (int -> tactic) -> int -> tactic) -> theory -> theory |
17 |
val addWrapper: (Proof.context -> (int -> tactic) -> int -> tactic) -> theory -> theory |
|
20289 | 18 |
val Swrap: Proof.context -> (int -> tactic) -> int -> tactic |
19 |
val wrap: Proof.context -> (int -> tactic) -> int -> tactic |
|
18728 | 20 |
val intro_bang: int option -> attribute |
21 |
val elim_bang: int option -> attribute |
|
22 |
val dest_bang: int option -> attribute |
|
23 |
val intro: int option -> attribute |
|
24 |
val elim: int option -> attribute |
|
25 |
val dest: int option -> attribute |
|
26 |
val intro_query: int option -> attribute |
|
27 |
val elim_query: int option -> attribute |
|
28 |
val dest_query: int option -> attribute |
|
29 |
val rule_del: attribute |
|
30529 | 30 |
val add: (int option -> attribute) -> (int option -> attribute) -> (int option -> attribute) -> |
31 |
attribute context_parser |
|
12350 | 32 |
end; |
33 |
||
33369 | 34 |
structure Context_Rules: CONTEXT_RULES = |
12350 | 35 |
struct |
36 |
||
37 |
||
38 |
(** rule declaration contexts **) |
|
39 |
||
18637 | 40 |
(* context data *) |
12350 | 41 |
|
33370 | 42 |
datatype rules = Rules of |
12350 | 43 |
{next: int, |
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
44 |
rules: (int * (Bires.kind * thm)) list, |
82812
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents:
82808
diff
changeset
|
45 |
netpairs: Bires.netpair list, |
51798 | 46 |
wrappers: |
47 |
((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list * |
|
48 |
((Proof.context -> (int -> tactic) -> int -> tactic) * stamp) list}; |
|
12350 | 49 |
|
50 |
fun make_rules next rules netpairs wrappers = |
|
51 |
Rules {next = next, rules = rules, netpairs = netpairs, wrappers = wrappers}; |
|
52 |
||
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
53 |
fun add_rule kind opt_weight th (Rules {next, rules, netpairs, wrappers}) = |
61049 | 54 |
let |
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
55 |
val weight = opt_weight |> \<^if_none>\<open>Bires.subgoals_of (Bires.kind_rule kind th)\<close>; |
82812
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents:
82808
diff
changeset
|
56 |
val tag = {weight = weight, index = next}; |
61049 | 57 |
val th' = Thm.trim_context th; |
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
58 |
val rules' = (weight, (kind, th')) :: rules; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
59 |
val netpairs' = netpairs |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
60 |
|> Bires.kind_map kind (Bires.insert_tagged_rule (tag, Bires.kind_rule kind th')); |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
61 |
in make_rules (next - 1) rules' netpairs' wrappers end; |
12350 | 62 |
|
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
39557
diff
changeset
|
63 |
fun del_rule0 th (rs as Rules {next, rules, netpairs, wrappers}) = |
12350 | 64 |
let |
22360
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents:
21506
diff
changeset
|
65 |
fun eq_th (_, (_, th')) = Thm.eq_thm_prop (th, th'); |
82807 | 66 |
fun del b netpair = Bires.delete_tagged_rule (b, th) netpair handle Net.DELETE => netpair; |
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
67 |
val rules' = filter_out eq_th rules; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
68 |
val netpairs' = map (del false o del true) netpairs; |
12350 | 69 |
in |
70 |
if not (exists eq_th rules) then rs |
|
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
71 |
else make_rules next rules' netpairs' wrappers |
12350 | 72 |
end; |
73 |
||
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
39557
diff
changeset
|
74 |
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
|
75 |
|
82816 | 76 |
structure Data = Generic_Data |
18637 | 77 |
( |
33370 | 78 |
type T = rules; |
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
79 |
val empty = make_rules ~1 [] Bires.kind_netpairs ([], []); |
33519 | 80 |
fun merge |
81 |
(Rules {rules = rules1, wrappers = (ws1, ws1'), ...}, |
|
12350 | 82 |
Rules {rules = rules2, wrappers = (ws2, ws2'), ...}) = |
83 |
let |
|
18637 | 84 |
val wrappers = |
18921 | 85 |
(Library.merge (eq_snd (op =)) (ws1, ws2), Library.merge (eq_snd (op =)) (ws1', ws2')); |
86 |
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
|
87 |
k1 = k2 andalso Thm.eq_thm_prop (th1, th2)) (rules1, rules2); |
12350 | 88 |
val next = ~ (length rules); |
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
89 |
val netpairs = |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
90 |
fold (fn (index, (weight, (kind, th))) => |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
91 |
Bires.kind_map kind |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
92 |
(Bires.insert_tagged_rule ({weight = weight, index = index}, (Bires.kind_elim kind, th)))) |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
93 |
(next upto ~1 ~~ rules) Bires.kind_netpairs; |
23227 | 94 |
in make_rules (next - 1) rules netpairs wrappers end; |
18637 | 95 |
); |
12350 | 96 |
|
22846 | 97 |
fun print_rules ctxt = |
98 |
let |
|
82816 | 99 |
val Rules {rules, ...} = Data.get (Context.Proof ctxt); |
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
100 |
fun prt_kind kind = |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
101 |
Pretty.big_list (Bires.kind_title kind ^ ":") |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
102 |
(map_filter (fn (_, (kind', th)) => |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
103 |
if kind = kind' 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
|
104 |
(sort (int_ord o apply2 fst) rules)); |
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
105 |
in Pretty.writeln (Pretty.chunks (map prt_kind Bires.kind_domain)) end; |
12350 | 106 |
|
107 |
||
108 |
(* access data *) |
|
109 |
||
82816 | 110 |
fun netpairs ctxt = let val Rules {netpairs, ...} = Data.get (Context.Proof ctxt) in netpairs end; |
12350 | 111 |
val netpair_bang = hd o netpairs; |
112 |
val netpair = hd o tl o netpairs; |
|
113 |
||
114 |
||
12399 | 115 |
(* retrieving rules *) |
116 |
||
61049 | 117 |
local |
118 |
||
82812
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents:
82808
diff
changeset
|
119 |
fun order weighted = |
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
wenzelm
parents:
82808
diff
changeset
|
120 |
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
|
121 |
|
12399 | 122 |
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
|
123 |
map snd (order weighted (Net.unify_term net t)); |
12399 | 124 |
|
125 |
fun find_erules _ [] = K [] |
|
12805 | 126 |
| find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact)); |
16424 | 127 |
|
12399 | 128 |
fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal); |
12380 | 129 |
|
61049 | 130 |
in |
12380 | 131 |
|
61049 | 132 |
fun find_rules_netpair ctxt weighted facts goal (inet, enet) = |
133 |
find_erules weighted facts enet @ find_irules weighted goal inet |
|
67649 | 134 |
|> map (Thm.transfer' ctxt); |
61049 | 135 |
|
136 |
fun find_rules ctxt weighted facts goal = |
|
137 |
map (find_rules_netpair ctxt weighted facts goal) (netpairs ctxt); |
|
138 |
||
139 |
end; |
|
12350 | 140 |
|
141 |
||
12399 | 142 |
(* wrappers *) |
143 |
||
18667 | 144 |
fun gen_add_wrapper upd w = |
82816 | 145 |
Context.theory_map (Data.map (fn Rules {next, rules, netpairs, wrappers} => |
18667 | 146 |
make_rules next rules netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers))); |
12350 | 147 |
|
148 |
val addSWrapper = gen_add_wrapper Library.apfst; |
|
149 |
val addWrapper = gen_add_wrapper Library.apsnd; |
|
150 |
||
151 |
||
152 |
fun gen_wrap which ctxt = |
|
82816 | 153 |
let val Rules {wrappers, ...} = Data.get (Context.Proof ctxt) |
51798 | 154 |
in fold_rev (fn (w, _) => w ctxt) (which wrappers) end; |
12350 | 155 |
|
156 |
val Swrap = gen_wrap #1; |
|
157 |
val wrap = gen_wrap #2; |
|
158 |
||
159 |
||
160 |
||
161 |
(** attributes **) |
|
162 |
||
163 |
(* add and del rules *) |
|
164 |
||
45375
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents:
39557
diff
changeset
|
165 |
|
82816 | 166 |
val rule_del = Thm.declaration_attribute (Data.map o del_rule); |
12350 | 167 |
|
18637 | 168 |
fun rule_add k view opt_w = |
82816 | 169 |
Thm.declaration_attribute (fn th => Data.map (add_rule k opt_w (view th) o del_rule th)); |
12350 | 170 |
|
82817
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
171 |
val intro_bang = rule_add Bires.intro_bang_kind I; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
172 |
val elim_bang = rule_add Bires.elim_bang_kind I; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
173 |
val dest_bang = rule_add Bires.elim_bang_kind Tactic.make_elim; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
174 |
val intro = rule_add Bires.intro_kind I; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
175 |
val elim = rule_add Bires.elim_kind I; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
176 |
val dest = rule_add Bires.elim_kind Tactic.make_elim; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
177 |
val intro_query = rule_add Bires.intro_query_kind I; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
178 |
val elim_query = rule_add Bires.elim_query_kind I; |
be1fb22d9e2a
clarified signature: more explicit type Bires.kind;
wenzelm
parents:
82816
diff
changeset
|
179 |
val dest_query = rule_add Bires.elim_query_kind Tactic.make_elim; |
12350 | 180 |
|
53171 | 181 |
val _ = Theory.setup |
182 |
(snd o Global_Theory.add_thms [((Binding.empty, Drule.equal_intr_rule), [intro_query NONE])]); |
|
18637 | 183 |
|
12350 | 184 |
|
18637 | 185 |
(* concrete syntax *) |
186 |
||
30529 | 187 |
fun add a b c x = |
27809
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
26463
diff
changeset
|
188 |
(Scan.lift ((Args.bang >> K a || Args.query >> K c || Scan.succeed b) -- |
36950 | 189 |
Scan.option Parse.nat) >> (fn (f, n) => f n)) x; |
18637 | 190 |
|
53171 | 191 |
val _ = Theory.setup |
67147 | 192 |
(Attrib.setup \<^binding>\<open>intro\<close> (add intro_bang intro intro_query) |
82818 | 193 |
"declaration of Pure introduction rule" #> |
67147 | 194 |
Attrib.setup \<^binding>\<open>elim\<close> (add elim_bang elim elim_query) |
82818 | 195 |
"declaration of Pure elimination rule" #> |
67147 | 196 |
Attrib.setup \<^binding>\<open>dest\<close> (add dest_bang dest dest_query) |
82818 | 197 |
"declaration of Pure destruction rule" #> |
67147 | 198 |
Attrib.setup \<^binding>\<open>rule\<close> (Scan.lift Args.del >> K rule_del) |
82818 | 199 |
"remove declaration of Pure intro/elim/dest rule"); |
12350 | 200 |
|
201 |
end; |