1 (* Title: Pure/Isar/method.ML
3 Author: Markus Wenzel, TU Muenchen
8 signature BASIC_METHOD =
10 val FINDGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq
11 val HEADGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq
13 val trace_rules: bool ref
14 val print_methods: theory -> unit
15 val Method: bstring -> (Args.src -> ProofContext.context -> method) -> string -> unit
21 val apply: method -> thm list -> cases_tactic
22 val RAW_METHOD_CASES: (thm list -> cases_tactic) -> method
23 val RAW_METHOD: (thm list -> tactic) -> method
24 val METHOD_CASES: (thm list -> cases_tactic) -> method
25 val METHOD: (thm list -> tactic) -> 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
35 val cheating: bool -> ProofContext.context -> method
36 val intro: thm list -> method
37 val elim: thm list -> method
38 val unfold: thm list -> ProofContext.context -> method
39 val fold: thm list -> ProofContext.context -> method
40 val atomize: bool -> method
42 val fact: thm list -> ProofContext.context -> method
43 val assumption: ProofContext.context -> method
44 val close: bool -> ProofContext.context -> method
45 val trace: ProofContext.context -> thm list -> unit
46 val rule_tac: thm list -> thm list -> int -> tactic
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
52 val iprover_tac: ProofContext.context -> int option -> int -> tactic
53 val bires_inst_tac: bool -> ProofContext.context -> (indexname * string) list ->
55 val set_tactic: (ProofContext.context -> thm list -> tactic) -> unit
56 val tactic: string -> ProofContext.context -> method
59 Basic of (ProofContext.context -> method) |
65 val primitive_text: (thm -> thm) -> text
66 val succeed_text: text
67 val default_text: text
70 val sorry_text: bool -> text
71 val finish_text: text option * bool -> text
72 exception METHOD_FAIL of (string * Position.T) * exn
73 val method: theory -> src -> ProofContext.context -> method
74 val add_methods: (bstring * (src -> ProofContext.context -> method) * string) list
76 val add_method: bstring * (src -> ProofContext.context -> method) * string
78 val method_setup: bstring * string * string -> theory -> theory
79 val syntax: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list))
80 -> src -> ProofContext.context -> ProofContext.context * 'a
81 val simple_args: (Args.T list -> 'a * Args.T list)
82 -> ('a -> ProofContext.context -> method) -> src -> ProofContext.context -> method
83 val ctxt_args: (ProofContext.context -> method) -> src -> ProofContext.context -> method
84 val no_args: method -> src -> ProofContext.context -> method
86 val sectioned_args: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
87 (Args.T list -> modifier * Args.T list) list ->
88 ('a -> ProofContext.context -> 'b) -> src -> ProofContext.context -> 'b
89 val bang_sectioned_args:
90 (Args.T list -> modifier * Args.T list) list ->
91 (thm list -> ProofContext.context -> 'a) -> src -> ProofContext.context -> 'a
92 val bang_sectioned_args':
93 (Args.T list -> modifier * Args.T list) list ->
94 (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
95 ('a -> thm list -> ProofContext.context -> 'b) -> src -> ProofContext.context -> 'b
96 val only_sectioned_args:
97 (Args.T list -> modifier * Args.T list) list ->
98 (ProofContext.context -> 'a) -> src -> ProofContext.context -> 'a
99 val thms_ctxt_args: (thm list -> ProofContext.context -> 'a) -> src ->
100 ProofContext.context -> 'a
101 val thms_args: (thm list -> 'a) -> src -> ProofContext.context -> 'a
102 val thm_args: (thm -> 'a) -> src -> ProofContext.context -> 'a
103 val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic)
104 -> src -> ProofContext.context -> method
105 val goal_args': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list))
106 -> ('a -> int -> tactic) -> src -> ProofContext.context -> method
107 val goal_args_ctxt: (Args.T list -> 'a * Args.T list) ->
108 (ProofContext.context -> 'a -> int -> tactic) -> src -> ProofContext.context -> method
109 val goal_args_ctxt': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
110 (ProofContext.context -> 'a -> int -> tactic) -> src -> ProofContext.context -> method
113 structure Method: METHOD =
116 (** generic tools **)
118 (* goal addressing *)
120 fun FINDGOAL tac st =
121 let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n)
122 in find 1 (Thm.nprems_of st) st end;
124 fun HEADGOAL tac = tac 1;
128 (** proof methods **)
130 (* datatype method *)
132 datatype method = Meth of thm list -> cases_tactic;
134 fun apply (Meth m) = m;
136 val RAW_METHOD_CASES = Meth;
138 fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac);
140 fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts =>
141 Seq.THEN (ALLGOALS Tactic.conjunction_tac, tac facts));
143 fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Tactic.conjunction_tac THEN tac facts);
145 val fail = METHOD (K no_tac);
146 val succeed = METHOD (K all_tac);
153 fun cut_rule_tac raw_rule =
155 val rule = Drule.forall_intr_vars raw_rule;
156 val revcut_rl = Drule.incr_indexes rule Drule.revcut_rl;
157 in Tactic.rtac (rule COMP revcut_rl) end;
161 fun insert_tac [] i = all_tac
162 | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts);
164 val insert_facts = METHOD (ALLGOALS o insert_tac);
165 fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms));
167 fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac);
168 fun SIMPLE_METHOD' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac));
173 (* shuffle subgoals *)
175 fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1)));
176 fun defer opt_i = METHOD (K (Tactic.defer_tac (the_default 1 opt_i)));
181 fun cheating int ctxt = METHOD (K (setmp quick_and_dirty (int orelse ! quick_and_dirty)
182 (SkipProof.cheat_tac (ProofContext.theory_of ctxt))));
185 (* unfold intro/elim rules *)
187 fun intro ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths));
188 fun elim ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths));
191 (* unfold/fold definitions *)
193 fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (LocalDefs.unfold_tac ctxt ths));
194 fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (LocalDefs.fold_tac ctxt ths));
197 (* atomize rule statements *)
199 fun atomize false = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o ObjectLogic.atomize_tac)
200 | atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac)));
203 (* this -- resolve facts directly *)
205 val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac));
208 (* fact -- composition by facts from context *)
210 fun fact [] ctxt = SIMPLE_METHOD' HEADGOAL (ProofContext.some_fact_tac ctxt)
211 | fact rules _ = SIMPLE_METHOD' HEADGOAL (ProofContext.fact_tac rules);
219 foldr (op APPEND') (K no_tac) (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths);
221 val refl_tac = SUBGOAL (fn (prop, i) =>
222 if can Logic.dest_equals (Logic.strip_assums_concl prop)
223 then Tactic.rtac Drule.reflexive_thm i else no_tac);
227 asm_tac (ProofContext.prems_of ctxt) APPEND'
230 fun assumption_tac ctxt [] = assm_tac ctxt
231 | assumption_tac _ [fact] = asm_tac [fact]
232 | assumption_tac _ _ = K no_tac;
236 fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt);
237 fun close immed ctxt = METHOD (K
238 (FILTER Thm.no_prems ((if immed then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac)));
243 (* rule etc. -- single-step refinements *)
245 val trace_rules = ref false;
247 fun trace ctxt rules =
248 conditional (! trace_rules andalso not (null rules)) (fn () =>
249 Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules)
250 |> Pretty.string_of |> tracing);
254 fun gen_rule_tac tac rules facts =
256 if null facts then tac rules i st
257 else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules))
258 THEN_ALL_NEW Tactic.norm_hhf_tac;
260 fun gen_arule_tac tac j rules facts =
261 EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac);
263 fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) =>
266 if not (null arg_rules) then arg_rules
267 else List.concat (ContextRules.find_rules false facts goal ctxt)
268 in trace ctxt rules; tac rules facts i end);
270 fun meth tac x = METHOD (HEADGOAL o tac x);
271 fun meth' tac x y = METHOD (HEADGOAL o tac x y);
275 val rule_tac = gen_rule_tac Tactic.resolve_tac;
276 val rule = meth rule_tac;
277 val some_rule_tac = gen_some_rule_tac rule_tac;
278 val some_rule = meth' some_rule_tac;
280 val erule = meth' (gen_arule_tac Tactic.eresolve_tac);
281 val drule = meth' (gen_arule_tac Tactic.dresolve_tac);
282 val frule = meth' (gen_arule_tac Tactic.forward_tac);
287 (* iprover -- intuitionistic proof search *)
291 val remdups_tac = SUBGOAL (fn (g, i) =>
292 let val prems = Logic.strip_assums_hyp g in
293 REPEAT_DETERM_N (length prems - length (gen_distinct op aconv prems))
294 (Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i)
297 fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac;
299 fun gen_eq_set e s1 s2 =
300 length s1 = length s2 andalso
301 gen_subset e (s1, s2) andalso gen_subset e (s2, s1);
303 val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist;
305 fun safe_step_tac ctxt =
306 ContextRules.Swrap ctxt
307 (eq_assume_tac ORELSE'
308 bires_tac true (ContextRules.netpair_bang ctxt));
310 fun unsafe_step_tac ctxt =
311 ContextRules.wrap ctxt
313 bires_tac false (ContextRules.netpair_bang ctxt) APPEND'
314 bires_tac false (ContextRules.netpair ctxt));
316 fun step_tac ctxt i =
317 REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE
318 REMDUPS (unsafe_step_tac ctxt) i;
320 fun intprover_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else
322 val ps = Logic.strip_assums_hyp g;
323 val c = Logic.strip_assums_concl g;
325 if member (fn ((ps1, c1), (ps2, c2)) =>
326 c1 aconv c2 andalso gen_eq_set op aconv ps1 ps2) gs (ps, c) then no_tac
327 else (step_tac ctxt THEN_ALL_NEW intprover_tac ctxt ((ps, c) :: gs) (d + 1) lim) i
332 fun iprover_tac ctxt opt_lim =
333 SELECT_GOAL (DEEPEN (2, the_default 20 opt_lim) (intprover_tac ctxt [] 0) 4 1);
338 (* rule_tac etc. -- refer to dynamic goal state!! *)
340 fun bires_inst_tac bires_flag ctxt insts thm =
342 val thy = ProofContext.theory_of ctxt;
343 (* Separate type and term insts *)
344 fun has_type_var ((x, _), _) = (case Symbol.explode x of
345 "'"::cs => true | cs => false);
346 val Tinsts = List.filter has_type_var insts;
347 val tinsts = filter_out has_type_var insts;
351 (* Preprocess state: extract environment information:
352 - variables and their types
353 - type variables and their sorts
354 - parameters and their types *)
355 val (types, sorts) = types_sorts st;
356 (* Process type insts: Tinsts_env *)
357 fun absent xi = error
358 ("No such variable in theorem: " ^ Syntax.string_of_vname xi);
359 val (rtypes, rsorts) = types_sorts thm;
361 let val S = case rsorts xi of SOME S => S | NONE => absent xi;
362 val T = Sign.read_typ (thy, sorts) s;
363 val U = TVar (xi, S);
364 in if Sign.typ_instance thy (T, U) then (U, T)
366 ("Instantiation of " ^ Syntax.string_of_vname xi ^ " fails")
368 val Tinsts_env = map readT Tinsts;
369 (* Preprocess rule: extract vars and their types, apply Tinsts *)
372 SOME T => typ_subst_atomic Tinsts_env T
373 | NONE => absent xi);
374 val (xis, ss) = Library.split_list tinsts;
375 val Ts = map get_typ xis;
376 val (_, _, Bi, _) = dest_state(st,i)
377 val params = Logic.strip_params Bi
378 (* params of subgoal i as string typ pairs *)
379 val params = rev(Term.rename_wrt_term Bi params)
380 (* as they are printed: bound variables with *)
381 (* the same name are renamed during printing *)
382 fun types' (a, ~1) = (case AList.lookup (op =) params a of
383 NONE => types (a, ~1)
385 | types' xi = types xi;
386 fun internal x = is_some (types' (x, ~1));
387 val used = Drule.add_used thm (Drule.add_used st []);
389 ProofContext.read_termTs_schematic ctxt internal types' sorts used (ss ~~ Ts);
390 val envT' = map (fn (ixn, T) =>
391 (TVar (ixn, the (rsorts ixn)), T)) envT @ Tinsts_env;
395 pairself (Thm.cterm_of thy) (Var (xi, fastype_of t), t))
397 (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2)
399 (* Lift and instantiate rule *)
400 val {maxidx, ...} = rep_thm st;
401 val paramTs = map #2 params
403 fun liftvar (Var ((a,j), T)) =
404 Var((a, j+inc), paramTs ---> Logic.incr_tvar inc T)
405 | liftvar t = raise TERM("Variable expected", [t]);
406 fun liftterm t = list_abs_free
407 (params, Logic.incr_indexes(paramTs,inc) t)
408 fun liftpair (cv,ct) =
409 (cterm_fun liftvar cv, cterm_fun liftterm ct)
410 val lifttvar = pairself (ctyp_of thy o Logic.incr_tvar inc);
411 val rule = Drule.instantiate
412 (map lifttvar envT', map liftpair cenv)
413 (Thm.lift_rule (Thm.cprem_of st i) thm)
415 if i > nprems_of st then no_tac st
417 compose_tac (bires_flag, rule, nprems_of thm) i
419 handle TERM (msg,_) => (warning msg; no_tac st)
420 | THM (msg,_,_) => (warning msg; no_tac st);
427 fun gen_inst _ tac _ (quant, ([], thms)) =
428 METHOD (fn facts => quant (insert_tac facts THEN' tac thms))
429 | gen_inst inst_tac _ ctxt (quant, (insts, [thm])) =
431 quant (insert_tac facts THEN' inst_tac ctxt insts thm))
432 | gen_inst _ _ _ _ = error "Cannot have instantiations with multiple rules";
434 (* Preserve Var indexes of rl; increment revcut_rl instead.
435 Copied from tactic.ML *)
436 fun make_elim_preserve rl =
437 let val {maxidx,...} = rep_thm rl
438 fun cvar xi = cterm_of ProtoPure.thy (Var(xi,propT));
440 instantiate ([], [(cvar("V",0), cvar("V",maxidx+1)),
441 (cvar("W",0), cvar("W",maxidx+1))]) revcut_rl
442 val arg = (false, rl, nprems_of rl)
443 val [th] = Seq.list_of (bicompose false arg 1 revcut_rl')
445 handle Bind => raise THM("make_elim_preserve", 1, [rl]);
449 val res_inst_meth = gen_inst (bires_inst_tac false) Tactic.resolve_tac;
451 val eres_inst_meth = gen_inst (bires_inst_tac true) Tactic.eresolve_tac;
455 (fn ctxt => fn insts => bires_inst_tac false ctxt insts o make_elim_preserve)
456 Tactic.cut_rules_tac;
460 (fn ctxt => fn insts => bires_inst_tac true ctxt insts o make_elim_preserve)
465 (fn ctxt => fn insts => fn rule =>
466 bires_inst_tac false ctxt insts (make_elim_preserve rule) THEN'
470 fun subgoal_tac ctxt sprop =
471 DETERM o bires_inst_tac false ctxt [(("psi", 0), sprop)] cut_rl;
473 fun subgoals_tac ctxt sprops = EVERY' (map (subgoal_tac ctxt) sprops);
475 fun thin_tac ctxt s =
476 bires_inst_tac true ctxt [(("V", 0), s)] thin_rl;
483 val tactic_ref = ref ((fn _ => raise Match): ProofContext.context -> thm list -> tactic);
484 fun set_tactic f = tactic_ref := f;
486 fun tactic txt ctxt = METHOD (fn facts =>
488 ("let fun tactic (ctxt: ProofContext.context) (facts: thm list) : tactic = \
489 \let val thm = ProofContext.get_thm_closure ctxt o PureThy.Name\n\
490 \ and thms = ProofContext.get_thms_closure ctxt o PureThy.Name in\n"
492 "\nend in Method.set_tactic tactic end")
494 Context.setmp (SOME (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts));
498 (** method syntax **)
505 Basic of (ProofContext.context -> method) |
508 Orelse of text list |
512 val primitive_text = Basic o K o SIMPLE_METHOD o PRIMITIVE;
513 val succeed_text = Basic (K succeed);
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));
517 val sorry_text = Basic o cheating;
519 fun finish_text (NONE, immed) = Basic (close immed)
520 | finish_text (SOME txt, immed) = Then [txt, Basic (close immed)];
523 (* method definitions *)
525 structure MethodsData = TheoryDataFun
527 val name = "Isar/methods";
528 type T = (((src -> ProofContext.context -> method) * string) * stamp) NameSpace.table;
530 val empty = NameSpace.empty_table;
533 fun merge _ tables = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUPS dups =>
534 error ("Attempt to merge different versions of method(s) " ^ commas_quote dups);
538 fun prt_meth (name, ((_, comment), _)) = Pretty.block
539 [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
541 [Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))]
542 |> Pretty.chunks |> Pretty.writeln
546 val _ = Context.add_setup MethodsData.init;
547 val print_methods = MethodsData.print;
552 exception METHOD_FAIL of (string * Position.T) * exn;
556 val (space, meths) = MethodsData.get thy;
559 val ((raw_name, _), pos) = Args.dest_src src;
560 val name = NameSpace.intern space raw_name;
562 (case Symtab.lookup meths name of
563 NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
564 | SOME ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src))
571 fun add_methods raw_meths thy =
573 val new_meths = raw_meths |> map (fn (name, f, comment) =>
574 (name, ((f, comment), stamp ())));
576 fun add meths = NameSpace.extend_table (Sign.naming_of thy) (meths, new_meths)
577 handle Symtab.DUPS dups =>
578 error ("Duplicate declaration of method(s) " ^ commas_quote dups);
579 in MethodsData.map add thy end;
581 val add_method = add_methods o Library.single;
583 fun Method name meth cmt = Context.>> (add_method (name, meth, cmt));
588 fun method_setup (name, txt, cmt) =
590 "val thm = PureThy.get_thm_closure (Context.the_context ()) o PureThy.Name;\n\
591 \val thms = PureThy.get_thms_closure (Context.the_context ()) o PureThy.Name;\n\
592 \val method: bstring * (Method.src -> ProofContext.context -> Proof.method) * string"
593 "Method.add_method method"
594 ("(" ^ Library.quote name ^ ", " ^ txt ^ ", " ^ Library.quote cmt ^ ")");
598 (** concrete syntax **)
602 fun syntax scan = Args.context_syntax "method" scan;
604 fun simple_args scan f src ctxt : method =
605 #2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt);
607 fun ctxt_args (f: ProofContext.context -> method) src ctxt =
608 #2 (syntax (Scan.succeed (f ctxt)) src ctxt);
610 fun no_args m = ctxt_args (K m);
615 type modifier = (ProofContext.context -> ProofContext.context) * attribute;
619 fun sect ss = Scan.first (map Scan.lift ss);
620 fun thms ss = Scan.repeat (Scan.unless (sect ss) Attrib.multi_thm) >> List.concat;
622 fun apply (f, att) (context, ths) = foldl_map att (Context.map_proof f context, ths);
624 fun section ss = (sect ss -- thms ss) :-- (fn (m, ths) => Scan.depend (fn context =>
625 Scan.succeed (apply m (context, ths)))) >> #2;
627 fun sectioned args ss = args -- Scan.repeat (section ss);
631 fun sectioned_args args ss f src ctxt =
632 let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt
635 fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f;
636 fun bang_sectioned_args' ss scan f =
637 sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f);
638 fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f);
640 fun thms_ctxt_args f = sectioned_args (thms []) [] f;
641 fun thms_args f = thms_ctxt_args (K o f);
642 fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected");
651 val introN = "intro";
656 fun modifier name kind kind' att =
657 Args.$$$ name |-- (kind >> K NONE || kind' |-- Args.nat --| Args.colon >> SOME)
658 >> (pair (I: ProofContext.context -> ProofContext.context) o att);
660 val iprover_modifiers =
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)];
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)));
682 fun nat_thms_args f = uncurry f oo
683 (#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.thms));
687 (Args.enum1 "and" (Scan.lift (Args.name -- (Args.$$$ "=" |-- Args.!!! Args.name))) --|
688 Scan.lift (Args.$$$ "in")) [] -- Attrib.thms;
690 fun inst_args f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts) src ctxt));
694 (Args.enum1 "and" (Scan.lift (Args.var -- (Args.$$$ "=" |-- Args.!!! Args.name))) --|
695 Scan.lift (Args.$$$ "in")) [] -- Attrib.thms;
697 fun inst_args_var f src ctxt =
698 f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts_var) src ctxt));
700 fun goal_args' args tac src ctxt = #2 (syntax (Args.goal_spec HEADGOAL -- args >>
701 (fn (quant, s) => SIMPLE_METHOD' quant (tac s))) src ctxt);
703 fun goal_args args tac = goal_args' (Scan.lift args) tac;
705 fun goal_args_ctxt' args tac src ctxt =
706 #2 (syntax (Args.goal_spec HEADGOAL -- args >>
707 (fn (quant, s) => SIMPLE_METHOD' quant (tac ctxt s))) src ctxt);
709 fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac;
712 (* misc tactic emulations *)
714 val subgoal_meth = goal_args_ctxt (Scan.repeat1 Args.name) subgoals_tac;
715 val thin_meth = goal_args_ctxt Args.name thin_tac;
716 val rename_meth = goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac;
717 val rotate_meth = goal_args (Scan.optional Args.int 1) Tactic.rotate_tac;
722 val _ = Context.add_setup (add_methods
723 [("fail", no_args fail, "force failure"),
724 ("succeed", no_args succeed, "succeed"),
725 ("-", no_args insert_facts, "do nothing (insert current facts only)"),
726 ("insert", thms_args insert, "insert theorems, ignoring facts (improper)"),
727 ("intro", thms_args intro, "repeatedly apply introduction rules"),
728 ("elim", thms_args elim, "repeatedly apply elimination rules"),
729 ("unfold", thms_ctxt_args unfold_meth, "unfold definitions"),
730 ("fold", thms_ctxt_args fold_meth, "fold definitions"),
731 ("atomize", (atomize o #2) oo syntax (Args.mode "full"),
732 "present local premises as object-level statements"),
733 ("iprover", iprover_meth, "intuitionistic proof search"),
734 ("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"),
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)"),
738 ("this", no_args this, "apply current facts as rules"),
739 ("fact", thms_ctxt_args fact, "composition by facts from context"),
740 ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
741 ("rule_tac", inst_args_var res_inst_meth, "apply rule (dynamic instantiation)"),
742 ("erule_tac", inst_args_var eres_inst_meth, "apply rule in elimination manner (dynamic instantiation)"),
743 ("drule_tac", inst_args_var dres_inst_meth, "apply rule in destruct manner (dynamic instantiation)"),
744 ("frule_tac", inst_args_var forw_inst_meth, "apply rule in forward manner (dynamic instantiation)"),
745 ("cut_tac", inst_args_var cut_inst_meth, "cut rule (dynamic instantiation)"),
746 ("subgoal_tac", subgoal_meth, "insert subgoal (dynamic instantiation)"),
747 ("thin_tac", thin_meth, "remove premise (dynamic instantiation)"),
748 ("rename_tac", rename_meth, "rename parameters of goal (dynamic instantiation)"),
749 ("rotate_tac", rotate_meth, "rotate assumptions of goal"),
750 ("tactic", simple_args Args.name tactic, "ML tactic as proof method")]);
753 (*final declarations of this structure!*)
754 val unfold = unfold_meth;
755 val fold = fold_meth;
759 structure BasicMethod: BASIC_METHOD = Method;