| author | huffman | 
| Mon, 07 Mar 2005 23:54:01 +0100 | |
| changeset 15587 | f363e6e080e7 | 
| parent 15574 | b1d1b5bfc464 | 
| child 15801 | d2f5ca3c048d | 
| 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 | val setup: (theory -> theory) list | 
| 53 | end; | |
| 54 | ||
| 55 | structure ContextRules: CONTEXT_RULES = | |
| 56 | struct | |
| 57 | ||
| 58 | ||
| 59 | (** rule declaration contexts **) | |
| 60 | ||
| 61 | (* rule kinds *) | |
| 62 | ||
| 63 | val intro_bangK = (0, false); | |
| 64 | val elim_bangK = (0, true); | |
| 65 | val introK = (1, false); | |
| 66 | val elimK = (1, true); | |
| 67 | val intro_queryK = (2, false); | |
| 68 | val elim_queryK = (2, true); | |
| 69 | ||
| 70 | val kind_names = | |
| 71 | [(intro_bangK, "safe introduction rules (intro!)"), | |
| 72 | (elim_bangK, "safe elimination rules (elim!)"), | |
| 73 | (introK, "introduction rules (intro)"), | |
| 74 | (elimK, "elimination rules (elim)"), | |
| 75 | (intro_queryK, "extra introduction rules (intro?)"), | |
| 76 | (elim_queryK, "extra elimination rules (elim?)")]; | |
| 77 | ||
| 78 | val rule_kinds = map #1 kind_names; | |
| 79 | val rule_indexes = distinct (map #1 rule_kinds); | |
| 80 | ||
| 81 | ||
| 82 | (* raw data *) | |
| 83 | ||
| 84 | type netpair = ((int * int) * (bool * thm)) Net.net * ((int * int) * (bool * thm)) Net.net; | |
| 85 | val empty_netpairs: netpair list = replicate (length rule_indexes) (Net.empty, Net.empty); | |
| 86 | ||
| 87 | datatype T = Rules of | |
| 88 |  {next: int,
 | |
| 89 | rules: (int * ((int * bool) * thm)) list, | |
| 90 | netpairs: netpair list, | |
| 91 | wrappers: (((int -> tactic) -> int -> tactic) * stamp) list * | |
| 92 | (((int -> tactic) -> int -> tactic) * stamp) list}; | |
| 93 | ||
| 94 | fun make_rules next rules netpairs wrappers = | |
| 95 |   Rules {next = next, rules = rules, netpairs = netpairs, wrappers = wrappers};
 | |
| 96 | ||
| 97 | fun add_rule (i, b) opt_w th (Rules {next, rules, netpairs, wrappers}) =
 | |
| 15531 | 98 | let val w = (case opt_w of SOME w => w | NONE => Tactic.subgoals_of_brl (b, th)) in | 
| 12350 | 99 | make_rules (next - 1) ((w, ((i, b), th)) :: rules) | 
| 100 | (map_nth_elem i (curry insert_tagged_brl ((w, next), (b, th))) netpairs) wrappers | |
| 101 | end; | |
| 102 | ||
| 103 | fun del_rule th (rs as Rules {next, rules, netpairs, wrappers}) =
 | |
| 104 | let | |
| 13105 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 wenzelm parents: 
12805diff
changeset | 105 | fun eq_th (_, (_, th')) = Drule.eq_thm_prop (th, th'); | 
| 12350 | 106 | fun del b netpair = delete_tagged_brl ((b, th), netpair) handle Net.DELETE => netpair; | 
| 107 | in | |
| 108 | if not (exists eq_th rules) then rs | |
| 109 | else make_rules next (filter_out eq_th rules) (map (del false o del true) netpairs) wrappers | |
| 110 | end; | |
| 111 | ||
| 112 | fun print_rules prt x (Rules {rules, ...}) =
 | |
| 113 | let | |
| 114 | fun prt_kind (i, b) = | |
| 15570 | 115 | Pretty.big_list (valOf (assoc (kind_names, (i, b))) ^ ":") | 
| 116 | (List.mapPartial (fn (_, (k, th)) => if k = (i, b) then SOME (prt x th) else NONE) | |
| 14472 
cba7c0a3ffb3
Removing the datatype declaration of "order" allows the standard General.order
 paulson parents: 
13372diff
changeset | 117 | (sort (Int.compare o pairself fst) rules)); | 
| 12350 | 118 | in Pretty.writeln (Pretty.chunks (map prt_kind rule_kinds)) end; | 
| 119 | ||
| 120 | ||
| 121 | (* theory and proof data *) | |
| 122 | ||
| 123 | structure GlobalRulesArgs = | |
| 124 | struct | |
| 125 | val name = "Isar/rule_context"; | |
| 126 | type T = T; | |
| 127 | ||
| 128 | val empty = make_rules ~1 [] empty_netpairs ([], []); | |
| 129 | val copy = I; | |
| 130 | val prep_ext = I; | |
| 131 | ||
| 132 |   fun merge (Rules {rules = rules1, wrappers = (ws1, ws1'), ...},
 | |
| 133 |       Rules {rules = rules2, wrappers = (ws2, ws2'), ...}) =
 | |
| 134 | let | |
| 135 | val wrappers = (gen_merge_lists' eq_snd ws1 ws2, gen_merge_lists' eq_snd ws1' ws2'); | |
| 136 | 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 | 137 | k1 = k2 andalso Drule.eq_thm_prop (th1, th2)) rules1 rules2; | 
| 12350 | 138 | val next = ~ (length rules); | 
| 15570 | 139 | val netpairs = Library.foldl (fn (nps, (n, (w, ((i, b), th)))) => | 
| 12350 | 140 | map_nth_elem i (curry insert_tagged_brl ((w, n), (b, th))) nps) | 
| 141 | (empty_netpairs, next upto ~1 ~~ rules); | |
| 142 | in make_rules (next - 1) rules netpairs wrappers end; | |
| 143 | ||
| 144 | val print = print_rules Display.pretty_thm_sg; | |
| 145 | end; | |
| 146 | ||
| 147 | structure GlobalRules = TheoryDataFun(GlobalRulesArgs); | |
| 148 | val print_global_rules = GlobalRules.print; | |
| 149 | ||
| 150 | structure LocalRulesArgs = | |
| 151 | struct | |
| 152 | val name = GlobalRulesArgs.name; | |
| 153 | type T = GlobalRulesArgs.T; | |
| 154 | val init = GlobalRules.get; | |
| 155 | val print = print_rules ProofContext.pretty_thm; | |
| 156 | end; | |
| 157 | ||
| 158 | structure LocalRules = ProofDataFun(LocalRulesArgs); | |
| 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 | ||
| 14472 
cba7c0a3ffb3
Removing the datatype declaration of "order" allows the standard General.order
 paulson parents: 
13372diff
changeset | 177 | fun orderlist brls = | 
| 
cba7c0a3ffb3
Removing the datatype declaration of "order" allows the standard General.order
 paulson parents: 
13372diff
changeset | 178 | untaglist (sort (prod_ord Int.compare Int.compare o pairself fst) brls); | 
| 
cba7c0a3ffb3
Removing the datatype declaration of "order" allows the standard General.order
 paulson parents: 
13372diff
changeset | 179 | fun orderlist_no_weight brls = | 
| 
cba7c0a3ffb3
Removing the datatype declaration of "order" allows the standard General.order
 paulson parents: 
13372diff
changeset | 180 | untaglist (sort (Int.compare o pairself (snd o fst)) brls); | 
| 12399 | 181 | |
| 182 | fun may_unify weighted t net = | |
| 183 | map snd ((if weighted then orderlist else orderlist_no_weight) (Net.unify_term net t)); | |
| 184 | ||
| 185 | fun find_erules _ [] = K [] | |
| 12805 | 186 | | find_erules w (fact :: _) = may_unify w (Logic.strip_assums_concl (Thm.prop_of fact)); | 
| 12399 | 187 | fun find_irules w goal = may_unify w (Logic.strip_assums_concl goal); | 
| 12380 | 188 | |
| 12399 | 189 | fun find_rules_netpair weighted facts goal (inet, enet) = | 
| 190 | find_erules weighted facts enet @ find_irules weighted goal inet; | |
| 12380 | 191 | |
| 12399 | 192 | fun find_rules weighted facts goals = map (find_rules_netpair weighted facts goals) o netpairs; | 
| 12350 | 193 | |
| 194 | ||
| 12399 | 195 | (* wrappers *) | 
| 196 | ||
| 12350 | 197 | fun gen_add_wrapper upd w = GlobalRules.map (fn (rs as Rules {next, rules, netpairs, wrappers}) =>
 | 
| 198 | make_rules next rules netpairs (upd (fn ws => (w, stamp ()) :: ws) wrappers)); | |
| 199 | ||
| 200 | val addSWrapper = gen_add_wrapper Library.apfst; | |
| 201 | val addWrapper = gen_add_wrapper Library.apsnd; | |
| 202 | ||
| 203 | ||
| 204 | fun gen_wrap which ctxt = | |
| 205 |   let val Rules {wrappers, ...} = LocalRules.get ctxt
 | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 206 | in fn tac => foldr (fn ((w, _), t) => w t) tac (which wrappers) end; | 
| 12350 | 207 | |
| 208 | val Swrap = gen_wrap #1; | |
| 209 | val wrap = gen_wrap #2; | |
| 210 | ||
| 211 | ||
| 212 | ||
| 213 | (** attributes **) | |
| 214 | ||
| 215 | (* add and del rules *) | |
| 216 | ||
| 217 | local | |
| 218 | ||
| 219 | fun del map_data (x, th) = | |
| 220 | (map_data (del_rule th o del_rule (Tactic.make_elim th)) x, th); | |
| 221 | ||
| 222 | fun add k view map_data opt_w = | |
| 12412 | 223 | (fn (x, th) => (map_data (add_rule k opt_w (view th)) x, th)) o del map_data; | 
| 12350 | 224 | |
| 225 | in | |
| 226 | ||
| 227 | val intro_bang_global = add intro_bangK I GlobalRules.map; | |
| 228 | val elim_bang_global = add elim_bangK I GlobalRules.map; | |
| 229 | val dest_bang_global = add elim_bangK Tactic.make_elim GlobalRules.map; | |
| 230 | val intro_global = add introK I GlobalRules.map; | |
| 231 | val elim_global = add elimK I GlobalRules.map; | |
| 232 | val dest_global = add elimK Tactic.make_elim GlobalRules.map; | |
| 233 | val intro_query_global = add intro_queryK I GlobalRules.map; | |
| 234 | val elim_query_global = add elim_queryK I GlobalRules.map; | |
| 235 | val dest_query_global = add elim_queryK Tactic.make_elim GlobalRules.map; | |
| 236 | val rule_del_global = del GlobalRules.map; | |
| 237 | ||
| 238 | val intro_bang_local = add intro_bangK I LocalRules.map; | |
| 239 | val elim_bang_local = add elim_bangK I LocalRules.map; | |
| 240 | val dest_bang_local = add elim_bangK Tactic.make_elim LocalRules.map; | |
| 241 | val intro_local = add introK I LocalRules.map; | |
| 242 | val elim_local = add elimK I LocalRules.map; | |
| 243 | val dest_local = add elimK Tactic.make_elim LocalRules.map; | |
| 244 | val intro_query_local = add intro_queryK I LocalRules.map; | |
| 245 | val elim_query_local = add elim_queryK I LocalRules.map; | |
| 246 | val dest_query_local = add elim_queryK Tactic.make_elim LocalRules.map; | |
| 247 | val rule_del_local = del LocalRules.map; | |
| 248 | ||
| 249 | end; | |
| 250 | ||
| 251 | ||
| 12380 | 252 | (* low-level modifiers *) | 
| 253 | ||
| 254 | fun modifier att (x, ths) = | |
| 255 | #1 (Thm.applys_attributes ((x, rev ths), [att])); | |
| 256 | ||
| 15531 | 257 | val addXIs_global = modifier (intro_query_global NONE); | 
| 258 | val addXEs_global = modifier (elim_query_global NONE); | |
| 259 | val addXDs_global = modifier (dest_query_global NONE); | |
| 12380 | 260 | val delrules_global = modifier rule_del_global; | 
| 261 | ||
| 15531 | 262 | val addXIs_local = modifier (intro_query_local NONE); | 
| 263 | val addXEs_local = modifier (elim_query_local NONE); | |
| 264 | val addXDs_local = modifier (dest_query_local NONE); | |
| 12380 | 265 | val delrules_local = modifier rule_del_local; | 
| 266 | ||
| 267 | ||
| 12350 | 268 | |
| 269 | (** theory setup **) | |
| 270 | ||
| 271 | val setup = | |
| 13372 
18790d503fe0
module now right after ProofContext (for locales);
 wenzelm parents: 
13105diff
changeset | 272 | [GlobalRules.init, LocalRules.init]; | 
| 12350 | 273 | |
| 274 | ||
| 275 | end; |