| author | paulson | 
| Tue, 10 Feb 2004 12:17:04 +0100 | |
| changeset 14379 | ea10a8c3e9cf | 
| parent 14215 | ebf291f3b449 | 
| child 14508 | 859b11514537 | 
| permissions | -rw-r--r-- | 
| 5824 | 1 | (* Title: Pure/Isar/method.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 8807 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 5824 | 5 | |
| 6 | Proof methods. | |
| 7 | *) | |
| 8 | ||
| 9 | signature BASIC_METHOD = | |
| 10 | sig | |
| 11731 | 11 | val trace_rules: bool ref | 
| 5824 | 12 | val print_methods: theory -> unit | 
| 13 | val Method: bstring -> (Args.src -> Proof.context -> Proof.method) -> string -> unit | |
| 14 | end; | |
| 15 | ||
| 16 | signature METHOD = | |
| 17 | sig | |
| 18 | include BASIC_METHOD | |
| 12055 | 19 | val trace: Proof.context -> thm list -> unit | 
| 12144 | 20 | val RAW_METHOD: (thm list -> tactic) -> Proof.method | 
| 21 | val RAW_METHOD_CASES: | |
| 22 | (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method | |
| 6091 | 23 | val METHOD: (thm list -> tactic) -> Proof.method | 
| 8372 | 24 | val METHOD_CASES: | 
| 25 | (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method | |
| 9706 | 26 | val SIMPLE_METHOD: tactic -> Proof.method | 
| 27 | val SIMPLE_METHOD': ((int -> tactic) -> tactic) -> (int -> tactic) -> Proof.method | |
| 5824 | 28 | val fail: Proof.method | 
| 29 | val succeed: Proof.method | |
| 8167 | 30 | val defer: int option -> Proof.method | 
| 31 | val prefer: int -> Proof.method | |
| 7419 | 32 | val insert_tac: thm list -> int -> tactic | 
| 7574 | 33 | val insert: thm list -> Proof.method | 
| 7555 | 34 | val insert_facts: Proof.method | 
| 7601 | 35 | val unfold: thm list -> Proof.method | 
| 7419 | 36 | val fold: thm list -> Proof.method | 
| 37 | val multi_resolve: thm list -> thm -> thm Seq.seq | |
| 38 | val multi_resolves: thm list -> thm list -> thm Seq.seq | |
| 12347 | 39 | val rules_tac: Proof.context -> int option -> int -> tactic | 
| 6091 | 40 | val rule_tac: thm list -> thm list -> int -> tactic | 
| 10309 | 41 | val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic | 
| 6091 | 42 | val rule: thm list -> Proof.method | 
| 10744 | 43 | val erule: int -> thm list -> Proof.method | 
| 44 | val drule: int -> thm list -> Proof.method | |
| 45 | val frule: int -> thm list -> Proof.method | |
| 8195 | 46 | val this: Proof.method | 
| 7555 | 47 | val assumption: Proof.context -> Proof.method | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 48 | val bires_inst_tac: bool -> Proof.context -> (indexname * string) list -> thm -> int -> tactic | 
| 8351 | 49 | val set_tactic: (Proof.context -> thm list -> tactic) -> unit | 
| 50 | val tactic: string -> Proof.context -> Proof.method | |
| 5916 | 51 | exception METHOD_FAIL of (string * Position.T) * exn | 
| 5824 | 52 | val method: theory -> Args.src -> Proof.context -> Proof.method | 
| 9539 | 53 | val add_method: bstring * (Args.src -> Proof.context -> Proof.method) * string | 
| 54 | -> theory -> theory | |
| 5824 | 55 | val add_methods: (bstring * (Args.src -> Proof.context -> Proof.method) * string) list | 
| 56 | -> theory -> theory | |
| 5884 | 57 | val syntax: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> | 
| 8282 | 58 | Args.src -> Proof.context -> Proof.context * 'a | 
| 8351 | 59 | val simple_args: (Args.T list -> 'a * Args.T list) | 
| 60 |     -> ('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
 | |
| 7555 | 61 | val ctxt_args: (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method | 
| 5884 | 62 | val no_args: Proof.method -> Args.src -> Proof.context -> Proof.method | 
| 7268 | 63 | type modifier | 
| 7601 | 64 | val sectioned_args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> | 
| 7268 | 65 | (Args.T list -> modifier * Args.T list) list -> | 
| 9864 | 66 |     ('a -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b
 | 
| 7601 | 67 | val bang_sectioned_args: | 
| 68 | (Args.T list -> modifier * Args.T list) list -> | |
| 9864 | 69 | (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a | 
| 9777 | 70 | val bang_sectioned_args': | 
| 71 | (Args.T list -> modifier * Args.T list) list -> | |
| 72 | (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> | |
| 9864 | 73 |     ('a -> thm list -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b
 | 
| 7601 | 74 | val only_sectioned_args: | 
| 75 | (Args.T list -> modifier * Args.T list) list -> | |
| 9864 | 76 | (Proof.context -> 'a) -> Args.src -> Proof.context -> 'a | 
| 77 | val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a | |
| 78 | val thms_args: (thm list -> 'a) -> Args.src -> Proof.context -> 'a | |
| 79 | val thm_args: (thm -> 'a) -> Args.src -> Proof.context -> 'a | |
| 5824 | 80 | datatype text = | 
| 81 | Basic of (Proof.context -> Proof.method) | | |
| 82 | Source of Args.src | | |
| 83 | Then of text list | | |
| 84 | Orelse of text list | | |
| 85 | Try of text | | |
| 86 | Repeat1 of text | |
| 87 | val refine: text -> Proof.state -> Proof.state Seq.seq | |
| 8238 | 88 | val refine_end: text -> Proof.state -> Proof.state Seq.seq | 
| 5824 | 89 | val proof: text option -> Proof.state -> Proof.state Seq.seq | 
| 8966 | 90 | val local_qed: bool -> text option | 
| 12144 | 91 | -> (Proof.context -> string * (string * thm list) list -> unit) * | 
| 12055 | 92 | (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq | 
| 6981 | 93 | val local_terminal_proof: text * text option | 
| 12144 | 94 | -> (Proof.context -> string * (string * thm list) list -> unit) * | 
| 12055 | 95 | (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq | 
| 12144 | 96 | val local_default_proof: (Proof.context -> string * (string * thm list) list -> unit) * | 
| 12055 | 97 | (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq | 
| 12144 | 98 | val local_immediate_proof: (Proof.context -> string * (string * thm list) list -> unit) * | 
| 12055 | 99 | (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq | 
| 12144 | 100 | val local_done_proof: (Proof.context -> string * (string * thm list) list -> unit) * | 
| 12055 | 101 | (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq | 
| 8966 | 102 | val global_qed: bool -> text option | 
| 12244 | 103 | -> Proof.state -> theory * ((string * string) * (string * thm list) list) | 
| 6934 | 104 | val global_terminal_proof: text * text option | 
| 12244 | 105 | -> Proof.state -> theory * ((string * string) * (string * thm list) list) | 
| 106 | val global_default_proof: Proof.state -> theory * ((string * string) * (string * thm list) list) | |
| 107 | val global_immediate_proof: Proof.state -> | |
| 108 | theory * ((string * string) * (string * thm list) list) | |
| 109 | val global_done_proof: Proof.state -> theory * ((string * string) * (string * thm list) list) | |
| 9539 | 110 |   val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic)
 | 
| 111 | -> Args.src -> Proof.context -> Proof.method | |
| 112 | val goal_args': (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) | |
| 113 |     -> ('a -> int -> tactic) -> Args.src -> Proof.context -> Proof.method
 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 114 | val goal_args_ctxt: (Args.T list -> 'a * Args.T list) -> (Proof.context -> 'a -> int -> tactic) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 115 | -> Args.src -> Proof.context -> Proof.method | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 116 | val goal_args_ctxt': (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 117 | -> (Proof.context -> 'a -> int -> tactic) -> Args.src -> Proof.context -> Proof.method | 
| 5824 | 118 | val setup: (theory -> theory) list | 
| 119 | end; | |
| 120 | ||
| 121 | structure Method: METHOD = | |
| 122 | struct | |
| 123 | ||
| 124 | ||
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 125 | (** proof methods **) | 
| 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 126 | |
| 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 127 | (* tracing *) | 
| 11731 | 128 | |
| 129 | val trace_rules = ref false; | |
| 130 | ||
| 12055 | 131 | fun trace ctxt rules = | 
| 12262 | 132 | conditional (! trace_rules andalso not (null rules)) (fn () => | 
| 133 | Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules) | |
| 134 | |> Pretty.string_of |> tracing); | |
| 11731 | 135 | |
| 136 | ||
| 8372 | 137 | (* make methods *) | 
| 5824 | 138 | |
| 12144 | 139 | val RAW_METHOD = Proof.method; | 
| 140 | val RAW_METHOD_CASES = Proof.method_cases; | |
| 141 | ||
| 142 | fun METHOD m = Proof.method (fn facts => TRY Tactic.conjunction_tac THEN m facts); | |
| 143 | fun METHOD_CASES m = | |
| 144 | Proof.method_cases (fn facts => Seq.THEN (TRY Tactic.conjunction_tac, m facts)); | |
| 8372 | 145 | |
| 5824 | 146 | |
| 147 | (* primitive *) | |
| 148 | ||
| 149 | val fail = METHOD (K no_tac); | |
| 150 | val succeed = METHOD (K all_tac); | |
| 151 | ||
| 152 | ||
| 8167 | 153 | (* shuffle *) | 
| 154 | ||
| 8240 | 155 | fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); | 
| 8167 | 156 | fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1))); | 
| 157 | ||
| 158 | ||
| 7419 | 159 | (* insert *) | 
| 160 | ||
| 161 | local | |
| 5824 | 162 | |
| 6981 | 163 | fun cut_rule_tac raw_rule = | 
| 164 | let | |
| 165 | val rule = Drule.forall_intr_vars raw_rule; | |
| 166 | val revcut_rl = Drule.incr_indexes_wrt [] [] [] [rule] Drule.revcut_rl; | |
| 7555 | 167 | in Tactic.rtac (rule COMP revcut_rl) end; | 
| 6981 | 168 | |
| 7419 | 169 | in | 
| 5824 | 170 | |
| 7419 | 171 | fun insert_tac [] i = all_tac | 
| 172 | | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); | |
| 6981 | 173 | |
| 7555 | 174 | val insert_facts = METHOD (ALLGOALS o insert_tac); | 
| 7664 | 175 | fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); | 
| 7419 | 176 | |
| 9706 | 177 | fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); | 
| 178 | fun SIMPLE_METHOD' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); | |
| 179 | ||
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 180 | end; | 
| 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 181 | |
| 9706 | 182 | |
| 12384 | 183 | (* unfold/fold definitions *) | 
| 184 | ||
| 185 | fun unfold ths = SIMPLE_METHOD (CHANGED_PROP (rewrite_goals_tac ths)); | |
| 186 | fun fold ths = SIMPLE_METHOD (CHANGED_PROP (fold_goals_tac ths)); | |
| 6532 | 187 | |
| 12384 | 188 | |
| 12829 | 189 | (* atomize rule statements *) | 
| 190 | ||
| 191 | fun atomize false = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o ObjectLogic.atomize_tac) | |
| 192 | | atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac))); | |
| 193 | ||
| 194 | ||
| 12384 | 195 | (* unfold intro/elim rules *) | 
| 196 | ||
| 197 | fun intro ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths)); | |
| 198 | fun elim ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths)); | |
| 9484 | 199 | |
| 200 | ||
| 7419 | 201 | (* multi_resolve *) | 
| 202 | ||
| 203 | local | |
| 204 | ||
| 205 | fun res th i rule = | |
| 206 | Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty; | |
| 207 | ||
| 208 | fun multi_res _ [] rule = Seq.single rule | |
| 209 | | multi_res i (th :: ths) rule = Seq.flat (Seq.map (res th i) (multi_res (i + 1) ths rule)); | |
| 210 | ||
| 211 | in | |
| 212 | ||
| 213 | val multi_resolve = multi_res 1; | |
| 8372 | 214 | fun multi_resolves facts rules = Seq.flat (Seq.map (multi_resolve facts) (Seq.of_list rules)); | 
| 7419 | 215 | |
| 216 | end; | |
| 217 | ||
| 218 | ||
| 12347 | 219 | (* rules_tac *) | 
| 8372 | 220 | |
| 7419 | 221 | local | 
| 5824 | 222 | |
| 12359 | 223 | val remdups_tac = SUBGOAL (fn (g, i) => | 
| 224 | let val prems = Logic.strip_assums_hyp g in | |
| 225 | REPEAT_DETERM_N (length prems - length (gen_distinct op aconv prems)) | |
| 226 | (Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i) | |
| 227 | end); | |
| 12347 | 228 | |
| 229 | fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac; | |
| 230 | ||
| 231 | fun gen_eq_set e s1 s2 = | |
| 232 | length s1 = length s2 andalso | |
| 233 | gen_subset e (s1, s2) andalso gen_subset e (s2, s1); | |
| 234 | ||
| 12350 | 235 | val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist; | 
| 12347 | 236 | |
| 237 | fun safe_step_tac ctxt = | |
| 12350 | 238 | ContextRules.Swrap ctxt | 
| 239 | (eq_assume_tac ORELSE' | |
| 240 | bires_tac true (ContextRules.netpair_bang ctxt)); | |
| 12347 | 241 | |
| 242 | fun unsafe_step_tac ctxt = | |
| 12350 | 243 | ContextRules.wrap ctxt | 
| 244 | (assume_tac APPEND' | |
| 245 | bires_tac false (ContextRules.netpair_bang ctxt) APPEND' | |
| 246 | bires_tac false (ContextRules.netpair ctxt)); | |
| 12347 | 247 | |
| 248 | fun step_tac ctxt i = | |
| 249 | REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE | |
| 250 | REMDUPS (unsafe_step_tac ctxt) i; | |
| 251 | ||
| 252 | fun intpr_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else | |
| 253 | let | |
| 254 | val ps = Logic.strip_assums_hyp g; | |
| 255 | val c = Logic.strip_assums_concl g; | |
| 256 | in | |
| 257 | if gen_mem (fn ((ps1, c1), (ps2, c2)) => | |
| 258 | c1 aconv c2 andalso gen_eq_set op aconv ps1 ps2) ((ps, c), gs) then no_tac | |
| 259 | else (step_tac ctxt THEN_ALL_NEW intpr_tac ctxt ((ps, c) :: gs) (d + 1) lim) i | |
| 260 | end); | |
| 261 | ||
| 262 | in | |
| 263 | ||
| 12359 | 264 | fun rules_tac ctxt opt_lim = | 
| 265 | SELECT_GOAL (DEEPEN (2, if_none opt_lim 20) (intpr_tac ctxt [] 0) 4 1); | |
| 12347 | 266 | |
| 267 | end; | |
| 268 | ||
| 269 | ||
| 270 | (* rule_tac etc. *) | |
| 271 | ||
| 272 | local | |
| 273 | ||
| 10541 
fdec07d4f047
resolveq(_cases)_tac moved to HOL/Tools/induct_method.ML;
 wenzelm parents: 
10529diff
changeset | 274 | fun gen_rule_tac tac rules [] i st = tac rules i st | 
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 275 | | gen_rule_tac tac rules facts i st = | 
| 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 276 | Seq.flat (Seq.map (fn rule => (tac o single) rule i st) (multi_resolves facts rules)); | 
| 7130 | 277 | |
| 10744 | 278 | fun gen_arule_tac tac j rules facts = | 
| 279 | EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); | |
| 280 | ||
| 11785 | 281 | fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) => | 
| 282 | let | |
| 283 | val rules = | |
| 284 | if not (null arg_rules) then arg_rules | |
| 12399 | 285 | else flat (ContextRules.find_rules false facts goal ctxt) | 
| 12055 | 286 | in trace ctxt rules; tac rules facts i end); | 
| 10309 | 287 | |
| 10744 | 288 | fun meth tac x = METHOD (HEADGOAL o tac x); | 
| 289 | fun meth' tac x y = METHOD (HEADGOAL o tac x y); | |
| 8220 | 290 | |
| 7419 | 291 | in | 
| 292 | ||
| 10744 | 293 | val rule_tac = gen_rule_tac Tactic.resolve_tac; | 
| 294 | val rule = meth rule_tac; | |
| 295 | val some_rule_tac = gen_some_rule_tac rule_tac; | |
| 296 | val some_rule = meth' some_rule_tac; | |
| 297 | ||
| 298 | val erule = meth' (gen_arule_tac Tactic.eresolve_tac); | |
| 299 | val drule = meth' (gen_arule_tac Tactic.dresolve_tac); | |
| 300 | val frule = meth' (gen_arule_tac Tactic.forward_tac); | |
| 5824 | 301 | |
| 7419 | 302 | end; | 
| 303 | ||
| 304 | ||
| 8195 | 305 | (* this *) | 
| 306 | ||
| 8671 | 307 | val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); | 
| 8195 | 308 | |
| 309 | ||
| 310 | (* assumption *) | |
| 7555 | 311 | |
| 10378 
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
 wenzelm parents: 
10309diff
changeset | 312 | fun asm_tac ths = | 
| 
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
 wenzelm parents: 
10309diff
changeset | 313 | foldr (op APPEND') (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths, K no_tac); | 
| 
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
 wenzelm parents: 
10309diff
changeset | 314 | |
| 10405 | 315 | fun assm_tac ctxt = | 
| 316 | assume_tac APPEND' | |
| 317 | asm_tac (ProofContext.prems_of ctxt) APPEND' | |
| 318 | Tactic.rtac Drule.reflexive_thm; | |
| 7419 | 319 | |
| 7555 | 320 | fun assumption_tac ctxt [] = assm_tac ctxt | 
| 10378 
98c95ebf804f
assumption / finish: handle non-atomic assumptions from context as well;
 wenzelm parents: 
10309diff
changeset | 321 | | assumption_tac _ [fact] = asm_tac [fact] | 
| 7555 | 322 | | assumption_tac _ _ = K no_tac; | 
| 7419 | 323 | |
| 8671 | 324 | fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt); | 
| 7419 | 325 | |
| 326 | ||
| 9539 | 327 | (* res_inst_tac etc. *) | 
| 8238 | 328 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 329 | (* Reimplemented to support both static (Isar) and dynamic (proof state) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 330 | context. By Clemens Ballarin. *) | 
| 12119 | 331 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 332 | fun bires_inst_tac bires_flag ctxt insts thm = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 333 | let | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 334 | val sign = ProofContext.sign_of ctxt; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 335 | (* Separate type and term insts *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 336 | fun has_type_var ((x, _), _) = (case Symbol.explode x of | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 337 | "'"::cs => true | cs => false); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 338 | val Tinsts = filter has_type_var insts; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 339 | val tinsts = filter_out has_type_var insts; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 340 | (* Tactic *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 341 | fun tac i st = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 342 | let | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 343 | (* Preprocess state: extract environment information: | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 344 | - variables and their types | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 345 | - type variables and their sorts | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 346 | - parameters and their types *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 347 | val (types, sorts) = types_sorts st; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 348 | (* Process type insts: Tinsts_env *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 349 | fun absent xi = error | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 350 | 	  ("No such variable in theorem: " ^ Syntax.string_of_vname xi);
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 351 | val (rtypes, rsorts) = types_sorts thm; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 352 | fun readT (xi, s) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 353 | let val S = case rsorts xi of Some S => S | None => absent xi; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 354 | val T = Sign.read_typ (sign, sorts) s; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 355 | in if Sign.typ_instance sign (T, TVar (xi, S)) then (xi, T) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 356 | else error | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 357 |              ("Instantiation of " ^ Syntax.string_of_vname xi ^ " fails")
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 358 | end; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 359 | val Tinsts_env = map readT Tinsts; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 360 | (* Preprocess rule: extract vars and their types, apply Tinsts *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 361 | fun get_typ xi = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 362 | (case rtypes xi of | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 363 | Some T => typ_subst_TVars Tinsts_env T | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 364 | | None => absent xi); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 365 | val (xis, ss) = Library.split_list tinsts; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 366 | val Ts = map get_typ xis; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 367 | val (_, _, Bi, _) = dest_state(st,i) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 368 | val params = Logic.strip_params Bi | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 369 | (* params of subgoal i as string typ pairs *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 370 | val params = rev(Term.rename_wrt_term Bi params) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 371 | (* as they are printed *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 372 | fun types' (a, ~1) = (case assoc (params, a) of | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 373 | None => types (a, ~1) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 374 | | some => some) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 375 | | types' xi = types xi; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 376 | val used = add_term_tvarnames | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 377 | (prop_of st $ prop_of thm,[]) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 378 | val (ts, envT) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 379 | ProofContext.read_termTs_env (types', sorts, used) ctxt (ss ~~ Ts); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 380 | val cenvT = map (apsnd (Thm.ctyp_of sign)) (envT @ Tinsts_env); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 381 | val cenv = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 382 | map | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 383 | (fn (xi, t) => | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 384 | pairself (Thm.cterm_of sign) (Var (xi, fastype_of t), t)) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 385 | (gen_distinct | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 386 | (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 387 | (xis ~~ ts)); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 388 | (* Lift and instantiate rule *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 389 | 	val {maxidx, ...} = rep_thm st;
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 390 | val paramTs = map #2 params | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 391 | and inc = maxidx+1 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 392 | fun liftvar (Var ((a,j), T)) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 393 | Var((a, j+inc), paramTs ---> incr_tvar inc T) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 394 | 	  | liftvar t = raise TERM("Variable expected", [t]);
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 395 | fun liftterm t = list_abs_free | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 396 | (params, Logic.incr_indexes(paramTs,inc) t) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 397 | fun liftpair (cv,ct) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 398 | (cterm_fun liftvar cv, cterm_fun liftterm ct) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 399 | fun lifttvar((a,i),ctyp) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 400 | 	    let val {T,sign} = rep_ctyp ctyp
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 401 | in ((a,i+inc), ctyp_of sign (incr_tvar inc T)) end | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 402 | val rule = Drule.instantiate | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 403 | (map lifttvar cenvT, map liftpair cenv) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 404 | (lift_rule (st, i) thm) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 405 | in | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 406 | if i > nprems_of st then no_tac st | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 407 | else st |> | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 408 | compose_tac (bires_flag, rule, nprems_of thm) i | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 409 | end | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 410 | handle TERM (msg,_) => (warning msg; no_tac st) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 411 | | THM (msg,_,_) => (warning msg; no_tac st); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 412 | in | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 413 | tac | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 414 | end; | 
| 8238 | 415 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 416 | fun gen_inst _ tac _ (quant, ([], thms)) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 417 | METHOD (fn facts => quant (insert_tac facts THEN' tac thms)) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 418 | | gen_inst inst_tac _ ctxt (quant, (insts, [thm])) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 419 | METHOD (fn facts => | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 420 | quant (insert_tac facts THEN' inst_tac ctxt insts thm)) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 421 | | gen_inst _ _ _ _ = error "Cannot have instantiations with multiple rules"; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 422 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 423 | val res_inst_meth = gen_inst (bires_inst_tac false) Tactic.resolve_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 424 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 425 | val eres_inst_meth = gen_inst (bires_inst_tac true) Tactic.eresolve_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 426 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 427 | (* Preserve Var indexes of rl; increment revcut_rl instead. | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 428 | Copied from tactic.ML *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 429 | fun make_elim_preserve rl = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 430 |   let val {maxidx,...} = rep_thm rl
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 431 | fun cvar xi = cterm_of (Theory.sign_of ProtoPure.thy) (Var(xi,propT)); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 432 | val revcut_rl' = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 433 |           instantiate ([],  [(cvar("V",0), cvar("V",maxidx+1)),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 434 |                              (cvar("W",0), cvar("W",maxidx+1))]) revcut_rl
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 435 | val arg = (false, rl, nprems_of rl) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 436 | val [th] = Seq.list_of (bicompose false arg 1 revcut_rl') | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 437 | in th end | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 438 |   handle Bind => raise THM("make_elim_preserve", 1, [rl]);
 | 
| 8238 | 439 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 440 | val cut_inst_meth = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 441 | gen_inst | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 442 | (fn ctxt => fn insts => fn thm => | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 443 | bires_inst_tac false ctxt insts (make_elim_preserve thm)) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 444 | Tactic.cut_rules_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 445 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 446 | val dres_inst_meth = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 447 | gen_inst | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 448 | (fn ctxt => fn insts => fn rule => | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 449 | bires_inst_tac true ctxt insts (make_elim_preserve rule)) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 450 | Tactic.dresolve_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 451 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 452 | val forw_inst_meth = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 453 | gen_inst | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 454 | (fn ctxt => fn insts => fn rule => | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 455 | bires_inst_tac false ctxt insts (make_elim_preserve rule) THEN' | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 456 | assume_tac) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 457 | Tactic.forward_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 458 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 459 | fun subgoal_tac ctxt sprop = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 460 |   DETERM o bires_inst_tac false ctxt [(("psi", 0), sprop)] cut_rl THEN'
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 461 | SUBGOAL (fn (prop, _) => | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 462 | let val concl' = Logic.strip_assums_concl prop in | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 463 | if null (term_tvars concl') then () | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 464 | else warning "Type variables in new subgoal: add a type constraint?"; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 465 | all_tac | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 466 | end); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 467 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 468 | fun subgoals_tac ctxt sprops = EVERY' (map (subgoal_tac ctxt) sprops); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 469 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 470 | fun thin_tac ctxt s = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 471 |   bires_inst_tac true ctxt [(("V", 0), s)] thin_rl;
 | 
| 8238 | 472 | |
| 8329 | 473 | (* simple Prolog interpreter *) | 
| 474 | ||
| 475 | fun prolog_tac rules facts = | |
| 476 | DEPTH_SOLVE_1 (HEADGOAL (Tactic.assume_tac APPEND' Tactic.resolve_tac (facts @ rules))); | |
| 477 | ||
| 478 | val prolog = METHOD o prolog_tac; | |
| 479 | ||
| 480 | ||
| 8351 | 481 | (* ML tactics *) | 
| 482 | ||
| 483 | val tactic_ref = ref ((fn _ => raise Match): Proof.context -> thm list -> tactic); | |
| 484 | fun set_tactic f = tactic_ref := f; | |
| 485 | ||
| 486 | fun tactic txt ctxt = METHOD (fn facts => | |
| 9631 | 487 | (Context.use_mltext | 
| 488 |     ("let fun tactic (ctxt: PureIsar.Proof.context) (facts: thm list) : tactic = \
 | |
| 489 | \let val thm = PureIsar.ProofContext.get_thm_closure ctxt\n\ | |
| 490 | \ and thms = PureIsar.ProofContext.get_thms_closure ctxt in\n" | |
| 491 | ^ txt ^ | |
| 492 | "\nend in PureIsar.Method.set_tactic tactic end") | |
| 493 | false None; | |
| 494 | Context.setmp (Some (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts)); | |
| 8351 | 495 | |
| 496 | ||
| 5824 | 497 | |
| 498 | (** methods theory data **) | |
| 499 | ||
| 500 | (* data kind 'Isar/methods' *) | |
| 501 | ||
| 502 | structure MethodsDataArgs = | |
| 503 | struct | |
| 504 | val name = "Isar/methods"; | |
| 505 | type T = | |
| 506 |     {space: NameSpace.T,
 | |
| 507 | meths: (((Args.src -> Proof.context -> Proof.method) * string) * stamp) Symtab.table}; | |
| 508 | ||
| 509 |   val empty = {space = NameSpace.empty, meths = Symtab.empty};
 | |
| 6546 | 510 | val copy = I; | 
| 5824 | 511 | val prep_ext = I; | 
| 512 |   fun merge ({space = space1, meths = meths1}, {space = space2, meths = meths2}) =
 | |
| 513 |     {space = NameSpace.merge (space1, space2),
 | |
| 514 | meths = Symtab.merge eq_snd (meths1, meths2) handle Symtab.DUPS dups => | |
| 515 |         error ("Attempt to merge different versions of methods " ^ commas_quote dups)};
 | |
| 516 | ||
| 9222 | 517 |   fun print _ {space, meths} =
 | 
| 5824 | 518 | let | 
| 519 | fun prt_meth (name, ((_, comment), _)) = Pretty.block | |
| 6849 | 520 | [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; | 
| 5824 | 521 | in | 
| 8720 | 522 | [Pretty.big_list "methods:" (map prt_meth (NameSpace.cond_extern_table space meths))] | 
| 9222 | 523 | |> Pretty.chunks |> Pretty.writeln | 
| 5824 | 524 | end; | 
| 525 | end; | |
| 526 | ||
| 527 | structure MethodsData = TheoryDataFun(MethodsDataArgs); | |
| 528 | val print_methods = MethodsData.print; | |
| 7611 | 529 | |
| 5824 | 530 | |
| 531 | (* get methods *) | |
| 532 | ||
| 5916 | 533 | exception METHOD_FAIL of (string * Position.T) * exn; | 
| 534 | ||
| 5824 | 535 | fun method thy = | 
| 536 | let | |
| 537 |     val {space, meths} = MethodsData.get thy;
 | |
| 538 | ||
| 5884 | 539 | fun meth src = | 
| 540 | let | |
| 541 | val ((raw_name, _), pos) = Args.dest_src src; | |
| 542 | val name = NameSpace.intern space raw_name; | |
| 543 | in | |
| 5824 | 544 | (case Symtab.lookup (meths, name) of | 
| 545 |           None => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
 | |
| 5916 | 546 | | Some ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src)) | 
| 5824 | 547 | end; | 
| 548 | in meth end; | |
| 549 | ||
| 550 | ||
| 9194 | 551 | (* add_method(s) *) | 
| 5824 | 552 | |
| 553 | fun add_methods raw_meths thy = | |
| 554 | let | |
| 555 | val full = Sign.full_name (Theory.sign_of thy); | |
| 556 | val new_meths = | |
| 557 | map (fn (name, f, comment) => (full name, ((f, comment), stamp ()))) raw_meths; | |
| 558 | ||
| 559 |     val {space, meths} = MethodsData.get thy;
 | |
| 560 | val space' = NameSpace.extend (space, map fst new_meths); | |
| 561 | val meths' = Symtab.extend (meths, new_meths) handle Symtab.DUPS dups => | |
| 562 |       error ("Duplicate declaration of method(s) " ^ commas_quote dups);
 | |
| 563 | in | |
| 564 |     thy |> MethodsData.put {space = space', meths = meths'}
 | |
| 565 | end; | |
| 566 | ||
| 9194 | 567 | val add_method = add_methods o Library.single; | 
| 568 | ||
| 5824 | 569 | (*implicit version*) | 
| 570 | fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]); | |
| 571 | ||
| 572 | ||
| 5884 | 573 | |
| 574 | (** method syntax **) | |
| 5824 | 575 | |
| 5884 | 576 | (* basic *) | 
| 577 | ||
| 578 | fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) = | |
| 579 | Args.syntax "method" scan; | |
| 5824 | 580 | |
| 8351 | 581 | fun simple_args scan f src ctxt : Proof.method = | 
| 582 | #2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); | |
| 583 | ||
| 7555 | 584 | fun ctxt_args (f: Proof.context -> Proof.method) src ctxt = | 
| 8282 | 585 | #2 (syntax (Scan.succeed (f ctxt)) src ctxt); | 
| 7555 | 586 | |
| 587 | fun no_args m = ctxt_args (K m); | |
| 5884 | 588 | |
| 589 | ||
| 590 | (* sections *) | |
| 5824 | 591 | |
| 7268 | 592 | type modifier = (Proof.context -> Proof.context) * Proof.context attribute; | 
| 593 | ||
| 594 | local | |
| 595 | ||
| 8381 | 596 | fun sect ss = Scan.first (map Scan.lift ss); | 
| 5884 | 597 | fun thms ss = Scan.unless (sect ss) Attrib.local_thms; | 
| 598 | fun thmss ss = Scan.repeat (thms ss) >> flat; | |
| 599 | ||
| 7268 | 600 | fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]); | 
| 5824 | 601 | |
| 7268 | 602 | fun section ss = (sect ss -- thmss ss) :-- (fn (m, ths) => Scan.depend (fn ctxt => | 
| 603 | Scan.succeed (apply m (ctxt, ths)))) >> #2; | |
| 5884 | 604 | |
| 7601 | 605 | fun sectioned args ss = args -- Scan.repeat (section ss); | 
| 5884 | 606 | |
| 7268 | 607 | in | 
| 5824 | 608 | |
| 5884 | 609 | fun sectioned_args args ss f src ctxt = | 
| 8282 | 610 | let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt | 
| 5921 | 611 | in f x ctxt' end; | 
| 5884 | 612 | |
| 7601 | 613 | fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; | 
| 9777 | 614 | fun bang_sectioned_args' ss scan f = | 
| 615 | sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f); | |
| 7601 | 616 | fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); | 
| 7268 | 617 | |
| 8093 | 618 | fun thms_ctxt_args f = sectioned_args (thmss []) [] f; | 
| 619 | fun thms_args f = thms_ctxt_args (K o f); | |
| 9706 | 620 | fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected"); | 
| 5824 | 621 | |
| 7268 | 622 | end; | 
| 623 | ||
| 5824 | 624 | |
| 12347 | 625 | (* rules syntax *) | 
| 626 | ||
| 627 | local | |
| 628 | ||
| 629 | val introN = "intro"; | |
| 630 | val elimN = "elim"; | |
| 631 | val destN = "dest"; | |
| 632 | val ruleN = "rule"; | |
| 633 | ||
| 634 | fun modifier name kind kind' att = | |
| 635 | Args.$$$ name |-- (kind >> K None || kind' |-- Args.nat --| Args.colon >> Some) | |
| 636 | >> (pair (I: Proof.context -> Proof.context) o att); | |
| 637 | ||
| 638 | val rules_modifiers = | |
| 12384 | 639 | [modifier destN Args.bang_colon Args.bang ContextRules.dest_bang_local, | 
| 12350 | 640 | modifier destN Args.colon (Scan.succeed ()) ContextRules.dest_local, | 
| 641 | modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang_local, | |
| 642 | modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim_local, | |
| 643 | modifier introN Args.bang_colon Args.bang ContextRules.intro_bang_local, | |
| 644 | modifier introN Args.colon (Scan.succeed ()) ContextRules.intro_local, | |
| 645 | Args.del -- Args.colon >> K (I, ContextRules.rule_del_local)]; | |
| 12347 | 646 | |
| 647 | in | |
| 648 | ||
| 649 | fun rules_args m = bang_sectioned_args' rules_modifiers (Scan.lift (Scan.option Args.nat)) m; | |
| 650 | ||
| 651 | fun rules_meth n prems ctxt = METHOD (fn facts => | |
| 12350 | 652 | HEADGOAL (insert_tac (prems @ facts) THEN' ObjectLogic.atomize_tac THEN' rules_tac ctxt n)); | 
| 12347 | 653 | |
| 654 | end; | |
| 655 | ||
| 656 | ||
| 9539 | 657 | (* tactic syntax *) | 
| 8238 | 658 | |
| 10744 | 659 | fun nat_thms_args f = uncurry f oo | 
| 660 | (#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.local_thmss)); | |
| 661 | ||
| 8238 | 662 | val insts = | 
| 9539 | 663 | Scan.optional | 
| 9565 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 664 | (Args.enum1 "and" (Scan.lift (Args.name -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| | 
| 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 665 | Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss; | 
| 8238 | 666 | |
| 12119 | 667 | fun inst_args f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts) src ctxt)); | 
| 8537 | 668 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 669 | val insts_var = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 670 | Scan.optional | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 671 | (Args.enum1 "and" (Scan.lift (Args.var -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 672 | Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 673 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 674 | fun inst_args_var f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts_var) src ctxt)); | 
| 8537 | 675 | |
| 12119 | 676 | fun goal_args' args tac src ctxt = #2 (syntax (Args.goal_spec HEADGOAL -- args >> | 
| 14215 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 ballarin parents: 
14174diff
changeset | 677 | (fn (quant, s) => SIMPLE_METHOD' quant (tac s))) src ctxt); | 
| 8537 | 678 | |
| 9539 | 679 | fun goal_args args tac = goal_args' (Scan.lift args) tac; | 
| 8238 | 680 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 681 | fun goal_args_ctxt' args tac src ctxt = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 682 | #2 (syntax (Args.goal_spec HEADGOAL -- args >> | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 683 | (fn (quant, s) => SIMPLE_METHOD' quant (tac ctxt s))) src ctxt); | 
| 8238 | 684 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 685 | fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac; | 
| 5824 | 686 | |
| 687 | (** method text **) | |
| 688 | ||
| 689 | (* datatype text *) | |
| 690 | ||
| 691 | datatype text = | |
| 692 | Basic of (Proof.context -> Proof.method) | | |
| 693 | Source of Args.src | | |
| 694 | Then of text list | | |
| 695 | Orelse of text list | | |
| 696 | Try of text | | |
| 697 | Repeat1 of text; | |
| 698 | ||
| 699 | ||
| 700 | (* refine *) | |
| 701 | ||
| 8238 | 702 | fun gen_refine f text state = | 
| 5824 | 703 | let | 
| 704 | val thy = Proof.theory_of state; | |
| 705 | ||
| 8238 | 706 | fun eval (Basic mth) = f mth | 
| 707 | | eval (Source src) = f (method thy src) | |
| 5824 | 708 | | eval (Then txts) = Seq.EVERY (map eval txts) | 
| 709 | | eval (Orelse txts) = Seq.FIRST (map eval txts) | |
| 710 | | eval (Try txt) = Seq.TRY (eval txt) | |
| 711 | | eval (Repeat1 txt) = Seq.REPEAT1 (eval txt); | |
| 712 | in eval text state end; | |
| 713 | ||
| 8238 | 714 | val refine = gen_refine Proof.refine; | 
| 715 | val refine_end = gen_refine Proof.refine_end; | |
| 6404 | 716 | |
| 5824 | 717 | |
| 6404 | 718 | (* structured proof steps *) | 
| 5824 | 719 | |
| 7506 | 720 | val default_text = Source (Args.src (("default", []), Position.none));
 | 
| 8195 | 721 | val this_text = Basic (K this); | 
| 9706 | 722 | val done_text = Basic (K (SIMPLE_METHOD all_tac)); | 
| 7555 | 723 | |
| 8966 | 724 | fun close_text asm = Basic (fn ctxt => METHOD (K | 
| 725 | (FILTER Thm.no_prems ((if asm then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac)))); | |
| 726 | ||
| 727 | fun finish_text asm None = close_text asm | |
| 728 | | finish_text asm (Some txt) = Then [txt, close_text asm]; | |
| 6872 | 729 | |
| 5824 | 730 | fun proof opt_text state = | 
| 731 | state | |
| 732 | |> Proof.assert_backward | |
| 6404 | 733 | |> refine (if_none opt_text default_text) | 
| 8242 | 734 | |> Seq.map (Proof.goal_facts (K [])) | 
| 5824 | 735 | |> Seq.map Proof.enter_forward; | 
| 736 | ||
| 8966 | 737 | fun local_qed asm opt_text = Proof.local_qed (refine (finish_text asm opt_text)); | 
| 738 | fun local_terminal_proof (text, opt_text) pr = | |
| 739 | Seq.THEN (proof (Some text), local_qed true opt_text pr); | |
| 740 | val local_default_proof = local_terminal_proof (default_text, None); | |
| 8195 | 741 | val local_immediate_proof = local_terminal_proof (this_text, None); | 
| 8966 | 742 | fun local_done_proof pr = Seq.THEN (proof (Some done_text), local_qed false None pr); | 
| 5824 | 743 | |
| 6872 | 744 | |
| 8966 | 745 | fun global_qeds asm opt_text = Proof.global_qed (refine (finish_text asm opt_text)); | 
| 5824 | 746 | |
| 8966 | 747 | fun global_qed asm opt_text state = | 
| 6872 | 748 | state | 
| 8966 | 749 | |> global_qeds asm opt_text | 
| 6872 | 750 | |> Proof.check_result "Failed to finish proof" state | 
| 751 | |> Seq.hd; | |
| 752 | ||
| 8966 | 753 | fun global_term_proof asm (text, opt_text) state = | 
| 6872 | 754 | state | 
| 755 | |> proof (Some text) | |
| 756 | |> Proof.check_result "Terminal proof method failed" state | |
| 8966 | 757 | |> (Seq.flat o Seq.map (global_qeds asm opt_text)) | 
| 6872 | 758 | |> Proof.check_result "Failed to finish proof (after successful terminal method)" state | 
| 759 | |> Seq.hd; | |
| 760 | ||
| 8966 | 761 | val global_terminal_proof = global_term_proof true; | 
| 6934 | 762 | val global_default_proof = global_terminal_proof (default_text, None); | 
| 8966 | 763 | val global_immediate_proof = global_terminal_proof (this_text, None); | 
| 764 | val global_done_proof = global_term_proof false (done_text, None); | |
| 5824 | 765 | |
| 766 | ||
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 767 | |
| 5824 | 768 | (** theory setup **) | 
| 769 | ||
| 9539 | 770 | (* misc tactic emulations *) | 
| 771 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 772 | val subgoal_meth = goal_args_ctxt (Scan.repeat1 Args.name) subgoals_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 773 | val thin_meth = goal_args_ctxt Args.name thin_tac; | 
| 9539 | 774 | val rename_meth = goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac; | 
| 9631 | 775 | val rotate_meth = goal_args (Scan.optional Args.int 1) Tactic.rotate_tac; | 
| 9539 | 776 | |
| 777 | ||
| 5824 | 778 | (* pure_methods *) | 
| 779 | ||
| 780 | val pure_methods = | |
| 781 |  [("fail", no_args fail, "force failure"),
 | |
| 782 |   ("succeed", no_args succeed, "succeed"),
 | |
| 9587 | 783 |   ("-", no_args insert_facts, "do nothing (insert current facts only)"),
 | 
| 9539 | 784 |   ("insert", thms_args insert, "insert theorems, ignoring facts (improper)"),
 | 
| 7601 | 785 |   ("unfold", thms_args unfold, "unfold definitions"),
 | 
| 12384 | 786 |   ("intro", thms_args intro, "repeatedly apply introduction rules"),
 | 
| 787 |   ("elim", thms_args elim, "repeatedly apply elimination rules"),
 | |
| 7601 | 788 |   ("fold", thms_args fold, "fold definitions"),
 | 
| 12829 | 789 |   ("atomize", (atomize o #2) oo syntax (Args.mode "full"),
 | 
| 11962 | 790 | "present local premises as object-level statements"), | 
| 12347 | 791 |   ("rules", rules_args rules_meth, "apply many rules, including proof search"),
 | 
| 12384 | 792 |   ("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"),
 | 
| 10744 | 793 |   ("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"),
 | 
| 794 |   ("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"),
 | |
| 795 |   ("frule", nat_thms_args frule, "apply rule in forward manner (improper)"),
 | |
| 8195 | 796 |   ("this", no_args this, "apply current facts as rules"),
 | 
| 8238 | 797 |   ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
 | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 798 |   ("rule_tac", inst_args_var res_inst_meth, "apply rule (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 799 |   ("erule_tac", inst_args_var eres_inst_meth, "apply rule in elimination manner (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 800 |   ("drule_tac", inst_args_var dres_inst_meth, "apply rule in destruct manner (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 801 |   ("frule_tac", inst_args_var forw_inst_meth, "apply rule in forward manner (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 802 |   ("cut_tac", inst_args_var cut_inst_meth, "cut rule (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 803 |   ("subgoal_tac", subgoal_meth, "insert subgoal (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 804 |   ("thin_tac", thin_meth, "remove premise (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 805 |   ("rename_tac", rename_meth, "rename parameters of goal (dynamic instantiation)"),
 | 
| 9631 | 806 |   ("rotate_tac", rotate_meth, "rotate assumptions of goal"),
 | 
| 8351 | 807 |   ("prolog", thms_args prolog, "simple prolog interpreter"),
 | 
| 808 |   ("tactic", simple_args Args.name tactic, "ML tactic as proof method")];
 | |
| 5824 | 809 | |
| 810 | ||
| 811 | (* setup *) | |
| 812 | ||
| 8153 
9bdbcb71dc56
maintain standard rules (beware: classical provers provides another version!);
 wenzelm parents: 
8093diff
changeset | 813 | val setup = | 
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 814 | [MethodsData.init, add_methods pure_methods, | 
| 12350 | 815 |   (#1 o PureThy.add_thms [(("", Drule.equal_intr_rule), [ContextRules.intro_query_global None])])];
 | 
| 5824 | 816 | |
| 817 | ||
| 818 | end; | |
| 819 | ||
| 820 | ||
| 821 | structure BasicMethod: BASIC_METHOD = Method; | |
| 822 | open BasicMethod; |