| author | wenzelm | 
| Fri, 05 May 2006 21:59:43 +0200 | |
| changeset 19575 | 2d9940cd52d3 | 
| parent 19482 | 9f11af8f7ef9 | 
| child 19778 | f0a318495ca4 | 
| permissions | -rw-r--r-- | 
| 5824 | 1 | (* Title: Pure/Isar/method.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 17110 | 5 | Isar proof methods. | 
| 5824 | 6 | *) | 
| 7 | ||
| 8 | signature BASIC_METHOD = | |
| 9 | sig | |
| 17110 | 10 | val FINDGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq | 
| 11 | val HEADGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq | |
| 12 | type method | |
| 11731 | 13 | val trace_rules: bool ref | 
| 5824 | 14 | val print_methods: theory -> unit | 
| 17110 | 15 | val Method: bstring -> (Args.src -> ProofContext.context -> method) -> string -> unit | 
| 5824 | 16 | end; | 
| 17 | ||
| 18 | signature METHOD = | |
| 19 | sig | |
| 20 | include BASIC_METHOD | |
| 18227 | 21 | val apply: method -> thm list -> cases_tactic | 
| 22 | val RAW_METHOD_CASES: (thm list -> cases_tactic) -> method | |
| 17110 | 23 | val RAW_METHOD: (thm list -> tactic) -> method | 
| 18227 | 24 | val METHOD_CASES: (thm list -> cases_tactic) -> method | 
| 17110 | 25 | val METHOD: (thm list -> tactic) -> method | 
| 26 | val fail: method | |
| 27 | val succeed: method | |
| 28 | val insert_tac: thm list -> int -> tactic | |
| 29 | val insert: thm list -> method | |
| 30 | val insert_facts: method | |
| 31 | val SIMPLE_METHOD: tactic -> method | |
| 32 | val SIMPLE_METHOD': ((int -> tactic) -> tactic) -> (int -> tactic) -> method | |
| 33 | val defer: int option -> method | |
| 34 | val prefer: int -> method | |
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 35 | val cheating: bool -> ProofContext.context -> method | 
| 17110 | 36 | val intro: thm list -> method | 
| 37 | val elim: thm list -> method | |
| 18877 | 38 | val unfold: thm list -> ProofContext.context -> method | 
| 39 | val fold: thm list -> ProofContext.context -> method | |
| 17110 | 40 | val atomize: bool -> method | 
| 41 | val this: method | |
| 18039 | 42 | val fact: thm list -> ProofContext.context -> method | 
| 17110 | 43 | val assumption: ProofContext.context -> method | 
| 44 | val close: bool -> ProofContext.context -> method | |
| 45 | val trace: ProofContext.context -> thm list -> unit | |
| 6091 | 46 | val rule_tac: thm list -> thm list -> int -> tactic | 
| 17110 | 47 | val some_rule_tac: thm list -> ProofContext.context -> thm list -> int -> tactic | 
| 48 | val rule: thm list -> method | |
| 49 | val erule: int -> thm list -> method | |
| 50 | val drule: int -> thm list -> method | |
| 51 | val frule: int -> thm list -> method | |
| 17587 | 52 | val iprover_tac: ProofContext.context -> int option -> int -> tactic | 
| 17110 | 53 | val bires_inst_tac: bool -> ProofContext.context -> (indexname * string) list -> | 
| 54 | thm -> int -> tactic | |
| 55 | val set_tactic: (ProofContext.context -> thm list -> tactic) -> unit | |
| 56 | val tactic: string -> ProofContext.context -> method | |
| 57 | type src | |
| 5824 | 58 | datatype text = | 
| 17110 | 59 | Basic of (ProofContext.context -> method) | | 
| 15703 | 60 | Source of src | | 
| 5824 | 61 | Then of text list | | 
| 62 | Orelse of text list | | |
| 63 | Try of text | | |
| 19186 | 64 | Repeat1 of text | | 
| 65 | SelectGoals of int * text | |
| 17857 | 66 | val primitive_text: (thm -> thm) -> text | 
| 67 | val succeed_text: text | |
| 17110 | 68 | val default_text: text | 
| 69 | val this_text: text | |
| 70 | val done_text: text | |
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 71 | val sorry_text: bool -> text | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 72 | val finish_text: text option * bool -> text | 
| 17110 | 73 | exception METHOD_FAIL of (string * Position.T) * exn | 
| 74 | val method: theory -> src -> ProofContext.context -> method | |
| 75 | val add_methods: (bstring * (src -> ProofContext.context -> method) * string) list | |
| 76 | -> theory -> theory | |
| 77 | val add_method: bstring * (src -> ProofContext.context -> method) * string | |
| 78 | -> theory -> theory | |
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 79 | val method_setup: bstring * string * string -> theory -> theory | 
| 18999 | 80 | val syntax: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) | 
| 17110 | 81 | -> src -> ProofContext.context -> ProofContext.context * 'a | 
| 82 | val simple_args: (Args.T list -> 'a * Args.T list) | |
| 83 |     -> ('a -> ProofContext.context -> method) -> src -> ProofContext.context -> method
 | |
| 84 | val ctxt_args: (ProofContext.context -> method) -> src -> ProofContext.context -> method | |
| 85 | val no_args: method -> src -> ProofContext.context -> method | |
| 86 | type modifier | |
| 18999 | 87 | val sectioned_args: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> | 
| 17110 | 88 | (Args.T list -> modifier * Args.T list) list -> | 
| 89 |     ('a -> ProofContext.context -> 'b) -> src -> ProofContext.context -> 'b
 | |
| 90 | val bang_sectioned_args: | |
| 91 | (Args.T list -> modifier * Args.T list) list -> | |
| 92 | (thm list -> ProofContext.context -> 'a) -> src -> ProofContext.context -> 'a | |
| 93 | val bang_sectioned_args': | |
| 94 | (Args.T list -> modifier * Args.T list) list -> | |
| 18999 | 95 | (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> | 
| 17110 | 96 |     ('a -> thm list -> ProofContext.context -> 'b) -> src -> ProofContext.context -> 'b
 | 
| 97 | val only_sectioned_args: | |
| 98 | (Args.T list -> modifier * Args.T list) list -> | |
| 99 | (ProofContext.context -> 'a) -> src -> ProofContext.context -> 'a | |
| 100 | val thms_ctxt_args: (thm list -> ProofContext.context -> 'a) -> src -> | |
| 101 | ProofContext.context -> 'a | |
| 102 | val thms_args: (thm list -> 'a) -> src -> ProofContext.context -> 'a | |
| 103 | val thm_args: (thm -> 'a) -> src -> ProofContext.context -> 'a | |
| 9539 | 104 |   val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic)
 | 
| 17110 | 105 | -> src -> ProofContext.context -> method | 
| 18999 | 106 | val goal_args': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) | 
| 17110 | 107 |     -> ('a -> int -> tactic) -> src -> ProofContext.context -> method
 | 
| 108 | val goal_args_ctxt: (Args.T list -> 'a * Args.T list) -> | |
| 109 | (ProofContext.context -> 'a -> int -> tactic) -> src -> ProofContext.context -> method | |
| 18999 | 110 | val goal_args_ctxt': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) -> | 
| 17110 | 111 | (ProofContext.context -> 'a -> int -> tactic) -> src -> ProofContext.context -> method | 
| 5824 | 112 | end; | 
| 113 | ||
| 114 | structure Method: METHOD = | |
| 115 | struct | |
| 116 | ||
| 17110 | 117 | (** generic tools **) | 
| 118 | ||
| 119 | (* goal addressing *) | |
| 120 | ||
| 121 | fun FINDGOAL tac st = | |
| 122 | let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n) | |
| 123 | in find 1 (Thm.nprems_of st) st end; | |
| 124 | ||
| 125 | fun HEADGOAL tac = tac 1; | |
| 126 | ||
| 127 | ||
| 5824 | 128 | |
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 129 | (** proof methods **) | 
| 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 130 | |
| 17110 | 131 | (* datatype method *) | 
| 11731 | 132 | |
| 18227 | 133 | datatype method = Meth of thm list -> cases_tactic; | 
| 11731 | 134 | |
| 17756 | 135 | fun apply (Meth m) = m; | 
| 11731 | 136 | |
| 17756 | 137 | val RAW_METHOD_CASES = Meth; | 
| 11731 | 138 | |
| 17110 | 139 | fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac); | 
| 12144 | 140 | |
| 17110 | 141 | fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts => | 
| 18474 | 142 | Seq.THEN (ALLGOALS Tactic.conjunction_tac, tac facts)); | 
| 8372 | 143 | |
| 18474 | 144 | fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Tactic.conjunction_tac THEN tac facts); | 
| 5824 | 145 | |
| 146 | val fail = METHOD (K no_tac); | |
| 147 | val succeed = METHOD (K all_tac); | |
| 148 | ||
| 149 | ||
| 17110 | 150 | (* insert facts *) | 
| 7419 | 151 | |
| 152 | local | |
| 5824 | 153 | |
| 6981 | 154 | fun cut_rule_tac raw_rule = | 
| 155 | let | |
| 156 | val rule = Drule.forall_intr_vars raw_rule; | |
| 18039 | 157 | val revcut_rl = Drule.incr_indexes rule Drule.revcut_rl; | 
| 7555 | 158 | in Tactic.rtac (rule COMP revcut_rl) end; | 
| 6981 | 159 | |
| 7419 | 160 | in | 
| 5824 | 161 | |
| 7419 | 162 | fun insert_tac [] i = all_tac | 
| 163 | | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts); | |
| 6981 | 164 | |
| 7555 | 165 | val insert_facts = METHOD (ALLGOALS o insert_tac); | 
| 7664 | 166 | fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms)); | 
| 7419 | 167 | |
| 9706 | 168 | fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac); | 
| 169 | fun SIMPLE_METHOD' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac)); | |
| 170 | ||
| 12324 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 171 | end; | 
| 
5db4b4596d1a
rule context and attributes moved to rule_context.ML;
 wenzelm parents: 
12311diff
changeset | 172 | |
| 9706 | 173 | |
| 17110 | 174 | (* shuffle subgoals *) | 
| 175 | ||
| 176 | fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1))); | |
| 18939 | 177 | fun defer opt_i = METHOD (K (Tactic.defer_tac (the_default 1 opt_i))); | 
| 17110 | 178 | |
| 179 | ||
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 180 | (* cheating *) | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 181 | |
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 182 | fun cheating int ctxt = METHOD (K (setmp quick_and_dirty (int orelse ! quick_and_dirty) | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 183 | (SkipProof.cheat_tac (ProofContext.theory_of ctxt)))); | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 184 | |
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 185 | |
| 17110 | 186 | (* unfold intro/elim rules *) | 
| 187 | ||
| 188 | fun intro ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths)); | |
| 189 | fun elim ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths)); | |
| 190 | ||
| 191 | ||
| 12384 | 192 | (* unfold/fold definitions *) | 
| 193 | ||
| 18877 | 194 | fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (LocalDefs.unfold_tac ctxt ths)); | 
| 195 | fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (LocalDefs.fold_tac ctxt ths)); | |
| 6532 | 196 | |
| 12384 | 197 | |
| 12829 | 198 | (* atomize rule statements *) | 
| 199 | ||
| 200 | fun atomize false = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o ObjectLogic.atomize_tac) | |
| 201 | | atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac))); | |
| 202 | ||
| 203 | ||
| 18039 | 204 | (* this -- resolve facts directly *) | 
| 12384 | 205 | |
| 17110 | 206 | val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac)); | 
| 9484 | 207 | |
| 208 | ||
| 18039 | 209 | (* fact -- composition by facts from context *) | 
| 210 | ||
| 18309 | 211 | fun fact [] ctxt = SIMPLE_METHOD' HEADGOAL (ProofContext.some_fact_tac ctxt) | 
| 212 | | fact rules _ = SIMPLE_METHOD' HEADGOAL (ProofContext.fact_tac rules); | |
| 18039 | 213 | |
| 214 | ||
| 17110 | 215 | (* assumption *) | 
| 7419 | 216 | |
| 217 | local | |
| 218 | ||
| 17110 | 219 | fun asm_tac ths = | 
| 220 | foldr (op APPEND') (K no_tac) (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths); | |
| 221 | ||
| 222 | val refl_tac = SUBGOAL (fn (prop, i) => | |
| 223 | if can Logic.dest_equals (Logic.strip_assums_concl prop) | |
| 224 | then Tactic.rtac Drule.reflexive_thm i else no_tac); | |
| 7419 | 225 | |
| 17110 | 226 | fun assm_tac ctxt = | 
| 227 | assume_tac APPEND' | |
| 228 | asm_tac (ProofContext.prems_of ctxt) APPEND' | |
| 229 | refl_tac; | |
| 230 | ||
| 231 | fun assumption_tac ctxt [] = assm_tac ctxt | |
| 232 | | assumption_tac _ [fact] = asm_tac [fact] | |
| 233 | | assumption_tac _ _ = K no_tac; | |
| 7419 | 234 | |
| 235 | in | |
| 236 | ||
| 17110 | 237 | fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt); | 
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 238 | fun close immed ctxt = METHOD (K | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 239 | (FILTER Thm.no_prems ((if immed then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac))); | 
| 7419 | 240 | |
| 241 | end; | |
| 242 | ||
| 243 | ||
| 17110 | 244 | (* rule etc. -- single-step refinements *) | 
| 12347 | 245 | |
| 17110 | 246 | val trace_rules = ref false; | 
| 12347 | 247 | |
| 17110 | 248 | fun trace ctxt rules = | 
| 249 | conditional (! trace_rules andalso not (null rules)) (fn () => | |
| 250 | Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules) | |
| 251 | |> Pretty.string_of |> tracing); | |
| 12347 | 252 | |
| 253 | local | |
| 254 | ||
| 18841 | 255 | fun gen_rule_tac tac rules facts = | 
| 256 | (fn i => fn st => | |
| 257 | if null facts then tac rules i st | |
| 258 | else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules)) | |
| 259 | THEN_ALL_NEW Tactic.norm_hhf_tac; | |
| 7130 | 260 | |
| 10744 | 261 | fun gen_arule_tac tac j rules facts = | 
| 262 | EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac); | |
| 263 | ||
| 11785 | 264 | fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) => | 
| 265 | let | |
| 266 | val rules = | |
| 267 | if not (null arg_rules) then arg_rules | |
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19307diff
changeset | 268 | else flat (ContextRules.find_rules false facts goal ctxt) | 
| 12055 | 269 | in trace ctxt rules; tac rules facts i end); | 
| 10309 | 270 | |
| 10744 | 271 | fun meth tac x = METHOD (HEADGOAL o tac x); | 
| 272 | fun meth' tac x y = METHOD (HEADGOAL o tac x y); | |
| 8220 | 273 | |
| 7419 | 274 | in | 
| 275 | ||
| 10744 | 276 | val rule_tac = gen_rule_tac Tactic.resolve_tac; | 
| 277 | val rule = meth rule_tac; | |
| 278 | val some_rule_tac = gen_some_rule_tac rule_tac; | |
| 279 | val some_rule = meth' some_rule_tac; | |
| 280 | ||
| 281 | val erule = meth' (gen_arule_tac Tactic.eresolve_tac); | |
| 282 | val drule = meth' (gen_arule_tac Tactic.dresolve_tac); | |
| 283 | val frule = meth' (gen_arule_tac Tactic.forward_tac); | |
| 5824 | 284 | |
| 7419 | 285 | end; | 
| 286 | ||
| 287 | ||
| 17587 | 288 | (* iprover -- intuitionistic proof search *) | 
| 17110 | 289 | |
| 290 | local | |
| 291 | ||
| 292 | val remdups_tac = SUBGOAL (fn (g, i) => | |
| 293 | let val prems = Logic.strip_assums_hyp g in | |
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
18999diff
changeset | 294 | REPEAT_DETERM_N (length prems - length (distinct op aconv prems)) | 
| 17110 | 295 | (Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i) | 
| 296 | end); | |
| 297 | ||
| 298 | fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac; | |
| 299 | ||
| 300 | val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist; | |
| 301 | ||
| 302 | fun safe_step_tac ctxt = | |
| 303 | ContextRules.Swrap ctxt | |
| 304 | (eq_assume_tac ORELSE' | |
| 305 | bires_tac true (ContextRules.netpair_bang ctxt)); | |
| 8195 | 306 | |
| 17110 | 307 | fun unsafe_step_tac ctxt = | 
| 308 | ContextRules.wrap ctxt | |
| 309 | (assume_tac APPEND' | |
| 310 | bires_tac false (ContextRules.netpair_bang ctxt) APPEND' | |
| 311 | bires_tac false (ContextRules.netpair ctxt)); | |
| 312 | ||
| 313 | fun step_tac ctxt i = | |
| 314 | REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE | |
| 315 | REMDUPS (unsafe_step_tac ctxt) i; | |
| 316 | ||
| 317 | fun intprover_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else | |
| 318 | let | |
| 319 | val ps = Logic.strip_assums_hyp g; | |
| 320 | val c = Logic.strip_assums_concl g; | |
| 321 | in | |
| 18921 | 322 | if member (fn ((ps1, c1), (ps2, c2)) => | 
| 19307 | 323 | c1 aconv c2 andalso | 
| 324 | length ps1 = length ps2 andalso | |
| 325 | gen_eq_set (op aconv) (ps1, ps2)) gs (ps, c) then no_tac | |
| 17110 | 326 | else (step_tac ctxt THEN_ALL_NEW intprover_tac ctxt ((ps, c) :: gs) (d + 1) lim) i | 
| 327 | end); | |
| 328 | ||
| 329 | in | |
| 330 | ||
| 17587 | 331 | fun iprover_tac ctxt opt_lim = | 
| 18939 | 332 | SELECT_GOAL (DEEPEN (2, the_default 20 opt_lim) (intprover_tac ctxt [] 0) 4 1); | 
| 17110 | 333 | |
| 334 | end; | |
| 8195 | 335 | |
| 336 | ||
| 17110 | 337 | (* rule_tac etc. -- refer to dynamic goal state!! *) | 
| 12119 | 338 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 339 | fun bires_inst_tac bires_flag ctxt insts thm = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 340 | let | 
| 17110 | 341 | val thy = ProofContext.theory_of ctxt; | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 342 | (* Separate type and term insts *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 343 | fun has_type_var ((x, _), _) = (case Symbol.explode x of | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 344 | "'"::cs => true | cs => false); | 
| 15570 | 345 | val Tinsts = List.filter has_type_var insts; | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 346 | val tinsts = filter_out has_type_var insts; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 347 | (* Tactic *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 348 | fun tac i st = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 349 | let | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 350 | (* Preprocess state: extract environment information: | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 351 | - variables and their types | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 352 | - type variables and their sorts | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 353 | - parameters and their types *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 354 | val (types, sorts) = types_sorts st; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 355 | (* Process type insts: Tinsts_env *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 356 | fun absent xi = error | 
| 14718 | 357 |           ("No such variable in theorem: " ^ Syntax.string_of_vname xi);
 | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 358 | val (rtypes, rsorts) = types_sorts thm; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 359 | fun readT (xi, s) = | 
| 15531 | 360 | let val S = case rsorts xi of SOME S => S | NONE => absent xi; | 
| 17110 | 361 | val T = Sign.read_typ (thy, sorts) s; | 
| 15798 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15713diff
changeset | 362 | val U = TVar (xi, S); | 
| 17110 | 363 | in if Sign.typ_instance thy (T, U) then (U, T) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 364 | else error | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 365 |              ("Instantiation of " ^ Syntax.string_of_vname xi ^ " fails")
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 366 | end; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 367 | val Tinsts_env = map readT Tinsts; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 368 | (* Preprocess rule: extract vars and their types, apply Tinsts *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 369 | fun get_typ xi = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 370 | (case rtypes xi of | 
| 15798 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15713diff
changeset | 371 | SOME T => typ_subst_atomic Tinsts_env T | 
| 15531 | 372 | | NONE => absent xi); | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 373 | val (xis, ss) = Library.split_list tinsts; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 374 | val Ts = map get_typ xis; | 
| 14718 | 375 | val (_, _, Bi, _) = dest_state(st,i) | 
| 376 | val params = Logic.strip_params Bi | |
| 377 | (* params of subgoal i as string typ pairs *) | |
| 378 | val params = rev(Term.rename_wrt_term Bi params) | |
| 379 | (* as they are printed: bound variables with *) | |
| 14508 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14215diff
changeset | 380 | (* the same name are renamed during printing *) | 
| 17314 | 381 | fun types' (a, ~1) = (case AList.lookup (op =) params a of | 
| 15531 | 382 | NONE => types (a, ~1) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 383 | | some => some) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 384 | | types' xi = types xi; | 
| 15973 | 385 | fun internal x = is_some (types' (x, ~1)); | 
| 15703 | 386 | val used = Drule.add_used thm (Drule.add_used st []); | 
| 14718 | 387 | val (ts, envT) = | 
| 388 | ProofContext.read_termTs_schematic ctxt internal types' sorts used (ss ~~ Ts); | |
| 15798 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15713diff
changeset | 389 | val envT' = map (fn (ixn, T) => | 
| 15973 | 390 | (TVar (ixn, the (rsorts ixn)), T)) envT @ Tinsts_env; | 
| 14718 | 391 | val cenv = | 
| 392 | map | |
| 393 | (fn (xi, t) => | |
| 17110 | 394 | pairself (Thm.cterm_of thy) (Var (xi, fastype_of t), t)) | 
| 19046 
bc5c6c9b114e
removed distinct, renamed gen_distinct to distinct;
 wenzelm parents: 
18999diff
changeset | 395 | (distinct | 
| 14718 | 396 | (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2) | 
| 397 | (xis ~~ ts)); | |
| 398 | (* Lift and instantiate rule *) | |
| 399 |         val {maxidx, ...} = rep_thm st;
 | |
| 400 | val paramTs = map #2 params | |
| 401 | and inc = maxidx+1 | |
| 402 | fun liftvar (Var ((a,j), T)) = | |
| 16876 | 403 | Var((a, j+inc), paramTs ---> Logic.incr_tvar inc T) | 
| 14718 | 404 |           | liftvar t = raise TERM("Variable expected", [t]);
 | 
| 405 | fun liftterm t = list_abs_free | |
| 406 | (params, Logic.incr_indexes(paramTs,inc) t) | |
| 407 | fun liftpair (cv,ct) = | |
| 408 | (cterm_fun liftvar cv, cterm_fun liftterm ct) | |
| 17110 | 409 | val lifttvar = pairself (ctyp_of thy o Logic.incr_tvar inc); | 
| 14718 | 410 | val rule = Drule.instantiate | 
| 15798 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 berghofe parents: 
15713diff
changeset | 411 | (map lifttvar envT', map liftpair cenv) | 
| 18145 | 412 | (Thm.lift_rule (Thm.cprem_of st i) thm) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 413 | in | 
| 14718 | 414 | if i > nprems_of st then no_tac st | 
| 415 | else st |> | |
| 416 | compose_tac (bires_flag, rule, nprems_of thm) i | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 417 | end | 
| 14718 | 418 | handle TERM (msg,_) => (warning msg; no_tac st) | 
| 419 | | THM (msg,_,_) => (warning msg; no_tac st); | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 420 | in | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 421 | tac | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 422 | end; | 
| 8238 | 423 | |
| 17110 | 424 | local | 
| 425 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 426 | fun gen_inst _ tac _ (quant, ([], thms)) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 427 | METHOD (fn facts => quant (insert_tac facts THEN' tac thms)) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 428 | | gen_inst inst_tac _ ctxt (quant, (insts, [thm])) = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 429 | METHOD (fn facts => | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 430 | quant (insert_tac facts THEN' inst_tac ctxt insts thm)) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 431 | | gen_inst _ _ _ _ = error "Cannot have instantiations with multiple rules"; | 
| 14718 | 432 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 433 | (* Preserve Var indexes of rl; increment revcut_rl instead. | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 434 | Copied from tactic.ML *) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 435 | fun make_elim_preserve rl = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 436 |   let val {maxidx,...} = rep_thm rl
 | 
| 17110 | 437 | fun cvar xi = cterm_of ProtoPure.thy (Var(xi,propT)); | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 438 | val revcut_rl' = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 439 |           instantiate ([],  [(cvar("V",0), cvar("V",maxidx+1)),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 440 |                              (cvar("W",0), cvar("W",maxidx+1))]) revcut_rl
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 441 | val arg = (false, rl, nprems_of rl) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 442 | val [th] = Seq.list_of (bicompose false arg 1 revcut_rl') | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 443 | in th end | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 444 |   handle Bind => raise THM("make_elim_preserve", 1, [rl]);
 | 
| 8238 | 445 | |
| 17110 | 446 | in | 
| 447 | ||
| 448 | val res_inst_meth = gen_inst (bires_inst_tac false) Tactic.resolve_tac; | |
| 449 | ||
| 450 | val eres_inst_meth = gen_inst (bires_inst_tac true) Tactic.eresolve_tac; | |
| 451 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 452 | val cut_inst_meth = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 453 | gen_inst | 
| 17110 | 454 | (fn ctxt => fn insts => bires_inst_tac false ctxt insts o make_elim_preserve) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 455 | Tactic.cut_rules_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 456 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 457 | val dres_inst_meth = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 458 | gen_inst | 
| 17110 | 459 | (fn ctxt => fn insts => bires_inst_tac true ctxt insts o make_elim_preserve) | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 460 | Tactic.dresolve_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 461 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 462 | val forw_inst_meth = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 463 | gen_inst | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 464 | (fn ctxt => fn insts => fn rule => | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 465 | bires_inst_tac false ctxt insts (make_elim_preserve rule) THEN' | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 466 | assume_tac) | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 467 | Tactic.forward_tac; | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 468 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 469 | fun subgoal_tac ctxt sprop = | 
| 17110 | 470 |   DETERM o bires_inst_tac false ctxt [(("psi", 0), sprop)] cut_rl;
 | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 471 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 472 | fun subgoals_tac ctxt sprops = EVERY' (map (subgoal_tac ctxt) sprops); | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 473 | |
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 474 | fun thin_tac ctxt s = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 475 |   bires_inst_tac true ctxt [(("V", 0), s)] thin_rl;
 | 
| 8238 | 476 | |
| 17110 | 477 | end; | 
| 8329 | 478 | |
| 479 | ||
| 8351 | 480 | (* ML tactics *) | 
| 481 | ||
| 17110 | 482 | val tactic_ref = ref ((fn _ => raise Match): ProofContext.context -> thm list -> tactic); | 
| 8351 | 483 | fun set_tactic f = tactic_ref := f; | 
| 484 | ||
| 485 | fun tactic txt ctxt = METHOD (fn facts => | |
| 9631 | 486 | (Context.use_mltext | 
| 17110 | 487 |     ("let fun tactic (ctxt: ProofContext.context) (facts: thm list) : tactic = \
 | 
| 16500 
09d43301b195
refl_tac: avoid failure of unification, i.e. confusing trace msg;
 wenzelm parents: 
16448diff
changeset | 488 | \let val thm = ProofContext.get_thm_closure ctxt o PureThy.Name\n\ | 
| 
09d43301b195
refl_tac: avoid failure of unification, i.e. confusing trace msg;
 wenzelm parents: 
16448diff
changeset | 489 | \ and thms = ProofContext.get_thms_closure ctxt o PureThy.Name in\n" | 
| 
09d43301b195
refl_tac: avoid failure of unification, i.e. confusing trace msg;
 wenzelm parents: 
16448diff
changeset | 490 | ^ txt ^ | 
| 
09d43301b195
refl_tac: avoid failure of unification, i.e. confusing trace msg;
 wenzelm parents: 
16448diff
changeset | 491 | "\nend in Method.set_tactic tactic end") | 
| 15531 | 492 | false NONE; | 
| 493 | Context.setmp (SOME (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts)); | |
| 8351 | 494 | |
| 495 | ||
| 5824 | 496 | |
| 17110 | 497 | (** method syntax **) | 
| 498 | ||
| 499 | (* method text *) | |
| 500 | ||
| 501 | type src = Args.src; | |
| 5824 | 502 | |
| 17110 | 503 | datatype text = | 
| 504 | Basic of (ProofContext.context -> method) | | |
| 505 | Source of src | | |
| 506 | Then of text list | | |
| 507 | Orelse of text list | | |
| 508 | Try of text | | |
| 19186 | 509 | Repeat1 of text | | 
| 510 | SelectGoals of int * text; | |
| 17110 | 511 | |
| 17857 | 512 | val primitive_text = Basic o K o SIMPLE_METHOD o PRIMITIVE; | 
| 513 | val succeed_text = Basic (K succeed); | |
| 17110 | 514 | val default_text = Source (Args.src (("default", []), Position.none));
 | 
| 515 | val this_text = Basic (K this); | |
| 516 | val done_text = Basic (K (SIMPLE_METHOD all_tac)); | |
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 517 | val sorry_text = Basic o cheating; | 
| 17110 | 518 | |
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 519 | fun finish_text (NONE, immed) = Basic (close immed) | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 520 | | finish_text (SOME txt, immed) = Then [txt, Basic (close immed)]; | 
| 17110 | 521 | |
| 522 | ||
| 523 | (* method definitions *) | |
| 5824 | 524 | |
| 16448 | 525 | structure MethodsData = TheoryDataFun | 
| 526 | (struct | |
| 5824 | 527 | val name = "Isar/methods"; | 
| 17110 | 528 | type T = (((src -> ProofContext.context -> method) * string) * stamp) NameSpace.table; | 
| 5824 | 529 | |
| 16347 | 530 | val empty = NameSpace.empty_table; | 
| 6546 | 531 | val copy = I; | 
| 16448 | 532 | val extend = I; | 
| 17496 | 533 | fun merge _ tables = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUPS dups => | 
| 16347 | 534 |     error ("Attempt to merge different versions of method(s) " ^ commas_quote dups);
 | 
| 5824 | 535 | |
| 16347 | 536 | fun print _ meths = | 
| 5824 | 537 | let | 
| 538 | fun prt_meth (name, ((_, comment), _)) = Pretty.block | |
| 6849 | 539 | [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment]; | 
| 5824 | 540 | in | 
| 16347 | 541 | [Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))] | 
| 9222 | 542 | |> Pretty.chunks |> Pretty.writeln | 
| 5824 | 543 | end; | 
| 16448 | 544 | end); | 
| 5824 | 545 | |
| 18708 | 546 | val _ = Context.add_setup MethodsData.init; | 
| 5824 | 547 | val print_methods = MethodsData.print; | 
| 7611 | 548 | |
| 5824 | 549 | |
| 550 | (* get methods *) | |
| 551 | ||
| 5916 | 552 | exception METHOD_FAIL of (string * Position.T) * exn; | 
| 553 | ||
| 5824 | 554 | fun method thy = | 
| 555 | let | |
| 16347 | 556 | val (space, meths) = MethodsData.get thy; | 
| 5884 | 557 | fun meth src = | 
| 558 | let | |
| 559 | val ((raw_name, _), pos) = Args.dest_src src; | |
| 560 | val name = NameSpace.intern space raw_name; | |
| 561 | in | |
| 17412 | 562 | (case Symtab.lookup meths name of | 
| 15531 | 563 |           NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
 | 
| 564 | | SOME ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src)) | |
| 5824 | 565 | end; | 
| 566 | in meth end; | |
| 567 | ||
| 568 | ||
| 17110 | 569 | (* add method *) | 
| 5824 | 570 | |
| 571 | fun add_methods raw_meths thy = | |
| 572 | let | |
| 16145 | 573 | val new_meths = raw_meths |> map (fn (name, f, comment) => | 
| 16347 | 574 | (name, ((f, comment), stamp ()))); | 
| 5824 | 575 | |
| 17110 | 576 | fun add meths = NameSpace.extend_table (Sign.naming_of thy) (meths, new_meths) | 
| 16347 | 577 | handle Symtab.DUPS dups => | 
| 578 |         error ("Duplicate declaration of method(s) " ^ commas_quote dups);
 | |
| 579 | in MethodsData.map add thy end; | |
| 5824 | 580 | |
| 9194 | 581 | val add_method = add_methods o Library.single; | 
| 582 | ||
| 18728 | 583 | fun Method name meth cmt = Context.>> (add_method (name, meth, cmt)); | 
| 5824 | 584 | |
| 585 | ||
| 17356 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 586 | (* method_setup *) | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 587 | |
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 588 | fun method_setup (name, txt, cmt) = | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 589 | Context.use_let | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 590 | "val thm = PureThy.get_thm_closure (Context.the_context ()) o PureThy.Name;\n\ | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 591 | \val thms = PureThy.get_thms_closure (Context.the_context ()) o PureThy.Name;\n\ | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 592 | \val method: bstring * (Method.src -> ProofContext.context -> Proof.method) * string" | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 593 | "Method.add_method method" | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 594 |     ("(" ^ Library.quote name ^ ", " ^ txt ^ ", " ^ Library.quote cmt ^ ")");
 | 
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 595 | |
| 
09afdf37cdb3
added cheating, sorry_text (from skip_proofs.ML);
 wenzelm parents: 
17314diff
changeset | 596 | |
| 5884 | 597 | |
| 17110 | 598 | (** concrete syntax **) | 
| 5824 | 599 | |
| 5884 | 600 | (* basic *) | 
| 601 | ||
| 18999 | 602 | fun syntax scan = Args.context_syntax "method" scan; | 
| 5824 | 603 | |
| 17110 | 604 | fun simple_args scan f src ctxt : method = | 
| 8351 | 605 | #2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt); | 
| 606 | ||
| 17110 | 607 | fun ctxt_args (f: ProofContext.context -> method) src ctxt = | 
| 8282 | 608 | #2 (syntax (Scan.succeed (f ctxt)) src ctxt); | 
| 7555 | 609 | |
| 610 | fun no_args m = ctxt_args (K m); | |
| 5884 | 611 | |
| 612 | ||
| 613 | (* sections *) | |
| 5824 | 614 | |
| 18728 | 615 | type modifier = (ProofContext.context -> ProofContext.context) * attribute; | 
| 7268 | 616 | |
| 617 | local | |
| 618 | ||
| 8381 | 619 | fun sect ss = Scan.first (map Scan.lift ss); | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19307diff
changeset | 620 | fun thms ss = Scan.repeat (Scan.unless (sect ss) Attrib.multi_thm) >> flat; | 
| 5884 | 621 | |
| 19186 | 622 | fun app (f, att) (context, ths) = foldl_map att (Context.map_proof f context, ths); | 
| 5824 | 623 | |
| 18999 | 624 | fun section ss = (sect ss -- thms ss) :-- (fn (m, ths) => Scan.depend (fn context => | 
| 19186 | 625 | Scan.succeed (app m (context, ths)))) >> #2; | 
| 5884 | 626 | |
| 7601 | 627 | fun sectioned args ss = args -- Scan.repeat (section ss); | 
| 5884 | 628 | |
| 7268 | 629 | in | 
| 5824 | 630 | |
| 5884 | 631 | fun sectioned_args args ss f src ctxt = | 
| 8282 | 632 | let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt | 
| 5921 | 633 | in f x ctxt' end; | 
| 5884 | 634 | |
| 7601 | 635 | fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f; | 
| 9777 | 636 | fun bang_sectioned_args' ss scan f = | 
| 637 | sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f); | |
| 7601 | 638 | fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f); | 
| 7268 | 639 | |
| 18999 | 640 | fun thms_ctxt_args f = sectioned_args (thms []) [] f; | 
| 8093 | 641 | fun thms_args f = thms_ctxt_args (K o f); | 
| 9706 | 642 | fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected"); | 
| 5824 | 643 | |
| 7268 | 644 | end; | 
| 645 | ||
| 5824 | 646 | |
| 17587 | 647 | (* iprover syntax *) | 
| 12347 | 648 | |
| 649 | local | |
| 650 | ||
| 651 | val introN = "intro"; | |
| 652 | val elimN = "elim"; | |
| 653 | val destN = "dest"; | |
| 654 | val ruleN = "rule"; | |
| 655 | ||
| 656 | fun modifier name kind kind' att = | |
| 15531 | 657 | Args.$$$ name |-- (kind >> K NONE || kind' |-- Args.nat --| Args.colon >> SOME) | 
| 17110 | 658 | >> (pair (I: ProofContext.context -> ProofContext.context) o att); | 
| 12347 | 659 | |
| 17587 | 660 | val iprover_modifiers = | 
| 18728 | 661 | [modifier destN Args.bang_colon Args.bang ContextRules.dest_bang, | 
| 662 | modifier destN Args.colon (Scan.succeed ()) ContextRules.dest, | |
| 663 | modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang, | |
| 664 | modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim, | |
| 665 | modifier introN Args.bang_colon Args.bang ContextRules.intro_bang, | |
| 666 | modifier introN Args.colon (Scan.succeed ()) ContextRules.intro, | |
| 667 | Args.del -- Args.colon >> K (I, ContextRules.rule_del)]; | |
| 12347 | 668 | |
| 669 | in | |
| 670 | ||
| 18640 | 671 | val iprover_meth = | 
| 672 | bang_sectioned_args' iprover_modifiers (Scan.lift (Scan.option Args.nat)) | |
| 673 | (fn n => fn prems => fn ctxt => METHOD (fn facts => | |
| 674 | HEADGOAL (insert_tac (prems @ facts) THEN' | |
| 675 | ObjectLogic.atomize_tac THEN' iprover_tac ctxt n))); | |
| 12347 | 676 | |
| 677 | end; | |
| 678 | ||
| 679 | ||
| 9539 | 680 | (* tactic syntax *) | 
| 8238 | 681 | |
| 10744 | 682 | fun nat_thms_args f = uncurry f oo | 
| 18999 | 683 | (#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.thms)); | 
| 10744 | 684 | |
| 8238 | 685 | val insts = | 
| 9539 | 686 | Scan.optional | 
| 9565 
3eb2ea15cc69
res_inst: include non-inst versions with multiple thms;
 wenzelm parents: 
9539diff
changeset | 687 | (Args.enum1 "and" (Scan.lift (Args.name -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| | 
| 18999 | 688 | Scan.lift (Args.$$$ "in")) [] -- Attrib.thms; | 
| 8238 | 689 | |
| 12119 | 690 | fun inst_args f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts) src ctxt)); | 
| 8537 | 691 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 692 | val insts_var = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 693 | Scan.optional | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 694 | (Args.enum1 "and" (Scan.lift (Args.var -- (Args.$$$ "=" |-- Args.!!! Args.name))) --| | 
| 18999 | 695 | Scan.lift (Args.$$$ "in")) [] -- Attrib.thms; | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 696 | |
| 17110 | 697 | fun inst_args_var f src ctxt = | 
| 698 | f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts_var) src ctxt)); | |
| 8537 | 699 | |
| 12119 | 700 | 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 | 701 | (fn (quant, s) => SIMPLE_METHOD' quant (tac s))) src ctxt); | 
| 8537 | 702 | |
| 9539 | 703 | fun goal_args args tac = goal_args' (Scan.lift args) tac; | 
| 8238 | 704 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 705 | fun goal_args_ctxt' args tac src ctxt = | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 706 | #2 (syntax (Args.goal_spec HEADGOAL -- args >> | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 707 | (fn (quant, s) => SIMPLE_METHOD' quant (tac ctxt s))) src ctxt); | 
| 8238 | 708 | |
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 709 | fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac; | 
| 5824 | 710 | |
| 14718 | 711 | |
| 9539 | 712 | (* misc tactic emulations *) | 
| 713 | ||
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 714 | 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 | 715 | val thin_meth = goal_args_ctxt Args.name thin_tac; | 
| 9539 | 716 | val rename_meth = goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac; | 
| 9631 | 717 | val rotate_meth = goal_args (Scan.optional Args.int 1) Tactic.rotate_tac; | 
| 9539 | 718 | |
| 719 | ||
| 18708 | 720 | (* theory setup *) | 
| 5824 | 721 | |
| 18708 | 722 | val _ = Context.add_setup (add_methods | 
| 5824 | 723 |  [("fail", no_args fail, "force failure"),
 | 
| 724 |   ("succeed", no_args succeed, "succeed"),
 | |
| 9587 | 725 |   ("-", no_args insert_facts, "do nothing (insert current facts only)"),
 | 
| 9539 | 726 |   ("insert", thms_args insert, "insert theorems, ignoring facts (improper)"),
 | 
| 12384 | 727 |   ("intro", thms_args intro, "repeatedly apply introduction rules"),
 | 
| 728 |   ("elim", thms_args elim, "repeatedly apply elimination rules"),
 | |
| 18877 | 729 |   ("unfold", thms_ctxt_args unfold_meth, "unfold definitions"),
 | 
| 730 |   ("fold", thms_ctxt_args fold_meth, "fold definitions"),
 | |
| 12829 | 731 |   ("atomize", (atomize o #2) oo syntax (Args.mode "full"),
 | 
| 11962 | 732 | "present local premises as object-level statements"), | 
| 18640 | 733 |   ("iprover", iprover_meth, "intuitionistic proof search"),
 | 
| 12384 | 734 |   ("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"),
 | 
| 10744 | 735 |   ("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"),
 | 
| 736 |   ("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"),
 | |
| 737 |   ("frule", nat_thms_args frule, "apply rule in forward manner (improper)"),
 | |
| 8195 | 738 |   ("this", no_args this, "apply current facts as rules"),
 | 
| 18039 | 739 |   ("fact", thms_ctxt_args fact, "composition by facts from context"),
 | 
| 8238 | 740 |   ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
 | 
| 14174 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 741 |   ("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 | 742 |   ("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 | 743 |   ("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 | 744 |   ("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 | 745 |   ("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 | 746 |   ("subgoal_tac", subgoal_meth, "insert subgoal (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 747 |   ("thin_tac", thin_meth, "remove premise (dynamic instantiation)"),
 | 
| 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 ballarin parents: 
13650diff
changeset | 748 |   ("rename_tac", rename_meth, "rename parameters of goal (dynamic instantiation)"),
 | 
| 9631 | 749 |   ("rotate_tac", rotate_meth, "rotate assumptions of goal"),
 | 
| 18708 | 750 |   ("tactic", simple_args Args.name tactic, "ML tactic as proof method")]);
 | 
| 5824 | 751 | |
| 752 | ||
| 16145 | 753 | (*final declarations of this structure!*) | 
| 754 | val unfold = unfold_meth; | |
| 755 | val fold = fold_meth; | |
| 756 | ||
| 5824 | 757 | end; | 
| 758 | ||
| 759 | structure BasicMethod: BASIC_METHOD = Method; | |
| 760 | open BasicMethod; |