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