reorganized structure Goal vs. Tactic;
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 -> Proof.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) -> method
33 val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method
34 val defer: int option -> method
35 val prefer: int -> method
36 val cheating: bool -> Proof.context -> method
37 val intro: thm list -> method
38 val elim: thm list -> method
39 val unfold: thm list -> Proof.context -> method
40 val fold: thm list -> Proof.context -> method
41 val atomize: bool -> method
43 val fact: thm list -> Proof.context -> method
44 val assumption: Proof.context -> method
45 val close: bool -> Proof.context -> method
46 val trace: Proof.context -> thm list -> unit
47 val rule_tac: thm list -> thm list -> int -> tactic
48 val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic
49 val rule: thm list -> method
50 val erule: int -> thm list -> method
51 val drule: int -> thm list -> method
52 val frule: int -> thm list -> method
53 val iprover_tac: Proof.context -> int option -> int -> tactic
54 val set_tactic: (Proof.context -> thm list -> tactic) -> unit
55 val tactic: string -> Proof.context -> method
58 Basic of (Proof.context -> method) |
65 SelectGoals of int * text
66 val primitive_text: (thm -> thm) -> text
67 val succeed_text: text
68 val default_text: text
71 val sorry_text: bool -> text
72 val finish_text: text option * bool -> text
73 exception METHOD_FAIL of (string * Position.T) * exn
74 val method: theory -> src -> Proof.context -> method
75 val method_i: theory -> src -> Proof.context -> method
76 val add_methods: (bstring * (src -> Proof.context -> method) * string) list
78 val add_method: bstring * (src -> Proof.context -> method) * string
80 val method_setup: bstring * string * string -> theory -> theory
81 val syntax: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list))
82 -> src -> Proof.context -> Proof.context * 'a
83 val simple_args: (Args.T list -> 'a * Args.T list)
84 -> ('a -> Proof.context -> method) -> src -> Proof.context -> method
85 val ctxt_args: (Proof.context -> method) -> src -> Proof.context -> method
86 val no_args: method -> src -> Proof.context -> method
88 val sectioned_args: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
89 (Args.T list -> modifier * Args.T list) list ->
90 ('a -> Proof.context -> 'b) -> src -> Proof.context -> 'b
91 val bang_sectioned_args:
92 (Args.T list -> modifier * Args.T list) list ->
93 (thm list -> Proof.context -> 'a) -> src -> Proof.context -> 'a
94 val bang_sectioned_args':
95 (Args.T list -> modifier * Args.T list) list ->
96 (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
97 ('a -> thm list -> Proof.context -> 'b) -> src -> Proof.context -> 'b
98 val only_sectioned_args:
99 (Args.T list -> modifier * Args.T list) list ->
100 (Proof.context -> 'a) -> src -> Proof.context -> 'a
101 val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> src ->
103 val thms_args: (thm list -> 'a) -> src -> Proof.context -> 'a
104 val thm_args: (thm -> 'a) -> src -> Proof.context -> 'a
105 val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic)
106 -> src -> Proof.context -> method
107 val goal_args': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list))
108 -> ('a -> int -> tactic) -> src -> Proof.context -> method
109 val goal_args_ctxt: (Args.T list -> 'a * Args.T list) ->
110 (Proof.context -> 'a -> int -> tactic) -> src -> Proof.context -> method
111 val goal_args_ctxt': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
112 (Proof.context -> 'a -> int -> tactic) -> src -> Proof.context -> method
115 structure Method: METHOD =
118 (** generic tools **)
120 (* goal addressing *)
122 fun FINDGOAL tac st =
123 let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n)
124 in find 1 (Thm.nprems_of st) st end;
126 fun HEADGOAL tac = tac 1;
130 (** proof methods **)
132 (* datatype method *)
134 datatype method = Meth of thm list -> cases_tactic;
136 fun apply (Meth m) = m;
138 val RAW_METHOD_CASES = Meth;
140 fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac);
142 fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts =>
143 Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts));
145 fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Goal.conjunction_tac THEN tac facts);
147 val fail = METHOD (K no_tac);
148 val succeed = METHOD (K all_tac);
155 fun cut_rule_tac rule =
156 Tactic.rtac (Drule.forall_intr_vars rule COMP_INCR revcut_rl);
160 fun insert_tac [] i = all_tac
161 | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts);
163 val insert_facts = METHOD (ALLGOALS o insert_tac);
164 fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms));
166 fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac);
167 fun SIMPLE_METHOD'' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac));
168 val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL;
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' (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths));
188 fun elim ths = SIMPLE_METHOD' (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' (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' (ProofContext.some_fact_tac ctxt)
211 | fact rules _ = SIMPLE_METHOD' (ProofContext.fact_tac rules);
219 foldr (op APPEND') (K no_tac) (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths);
221 fun cond_rtac cond rule = SUBGOAL (fn (prop, i) =>
222 if cond (Logic.strip_assums_concl prop)
223 then Tactic.rtac rule i else no_tac);
227 asm_tac (Assumption.prems_of ctxt) APPEND'
228 cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND'
229 cond_rtac (can Logic.dest_term) Drule.termI;
231 fun assumption_tac ctxt [] = assm_tac ctxt
232 | assumption_tac _ [fact] = asm_tac [fact]
233 | assumption_tac _ _ = K no_tac;
237 fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt);
238 fun close immed ctxt = METHOD (K
239 (FILTER Thm.no_prems ((if immed then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac)));
244 (* rule etc. -- single-step refinements *)
246 val trace_rules = ref false;
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);
255 fun gen_rule_tac tac rules facts =
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 Goal.norm_hhf_tac;
261 fun gen_arule_tac tac j rules facts =
262 EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac);
264 fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) =>
267 if not (null arg_rules) then arg_rules
268 else flat (ContextRules.find_rules false facts goal ctxt)
269 in trace ctxt rules; tac rules facts i end);
271 fun meth tac x = METHOD (HEADGOAL o tac x);
272 fun meth' tac x y = METHOD (HEADGOAL o tac x y);
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;
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);
288 (* iprover -- intuitionistic proof search *)
292 val remdups_tac = SUBGOAL (fn (g, i) =>
293 let val prems = Logic.strip_assums_hyp g in
294 REPEAT_DETERM_N (length prems - length (distinct op aconv prems))
295 (Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i)
298 fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac;
300 val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist;
302 fun safe_step_tac ctxt =
303 ContextRules.Swrap ctxt
304 (eq_assume_tac ORELSE'
305 bires_tac true (ContextRules.netpair_bang ctxt));
307 fun unsafe_step_tac ctxt =
308 ContextRules.wrap ctxt
310 bires_tac false (ContextRules.netpair_bang ctxt) APPEND'
311 bires_tac false (ContextRules.netpair ctxt));
313 fun step_tac ctxt i =
314 REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE
315 REMDUPS (unsafe_step_tac ctxt) i;
317 fun intprover_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else
319 val ps = Logic.strip_assums_hyp g;
320 val c = Logic.strip_assums_concl g;
322 if member (fn ((ps1, c1), (ps2, c2)) =>
324 length ps1 = length ps2 andalso
325 gen_eq_set (op aconv) (ps1, ps2)) gs (ps, c) then no_tac
326 else (step_tac ctxt THEN_ALL_NEW intprover_tac ctxt ((ps, c) :: gs) (d + 1) lim) i
331 fun iprover_tac ctxt opt_lim =
332 SELECT_GOAL (DEEPEN (2, the_default 20 opt_lim) (intprover_tac ctxt [] 0) 4 1);
339 val tactic_ref = ref ((fn _ => raise Match): Proof.context -> thm list -> tactic);
340 fun set_tactic f = tactic_ref := f;
342 fun tactic txt ctxt = METHOD (fn facts =>
344 ("let fun tactic (ctxt: Proof.context) (facts: thm list) : tactic = \
345 \let val thm = ProofContext.get_thm_closure ctxt o PureThy.Name\n\
346 \ and thms = ProofContext.get_thms_closure ctxt o PureThy.Name in\n"
348 "\nend in Method.set_tactic tactic end")
350 Context.setmp (SOME (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts));
354 (** method syntax **)
361 Basic of (Proof.context -> method) |
365 Orelse of text list |
368 SelectGoals of int * text;
370 val primitive_text = Basic o K o SIMPLE_METHOD o PRIMITIVE;
371 val succeed_text = Basic (K succeed);
372 val default_text = Source (Args.src (("default", []), Position.none));
373 val this_text = Basic (K this);
374 val done_text = Basic (K (SIMPLE_METHOD all_tac));
375 val sorry_text = Basic o cheating;
377 fun finish_text (NONE, immed) = Basic (close immed)
378 | finish_text (SOME txt, immed) = Then [txt, Basic (close immed)];
381 (* method definitions *)
383 structure MethodsData = TheoryDataFun
385 val name = "Isar/methods";
386 type T = (((src -> Proof.context -> method) * string) * stamp) NameSpace.table;
388 val empty = NameSpace.empty_table;
391 fun merge _ tables = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUPS dups =>
392 error ("Attempt to merge different versions of method(s) " ^ commas_quote dups);
396 fun prt_meth (name, ((_, comment), _)) = Pretty.block
397 [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
399 [Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))]
400 |> Pretty.chunks |> Pretty.writeln
404 val _ = Context.add_setup MethodsData.init;
405 val print_methods = MethodsData.print;
410 exception METHOD_FAIL of (string * Position.T) * exn;
414 val meths = #2 (MethodsData.get thy);
416 let val ((name, _), pos) = Args.dest_src src in
417 (case Symtab.lookup meths name of
418 NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
419 | SOME ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src))
423 fun method thy = method_i thy o Args.map_name (NameSpace.intern (#1 (MethodsData.get thy)));
428 fun add_methods raw_meths thy =
430 val new_meths = raw_meths |> map (fn (name, f, comment) =>
431 (name, ((f, comment), stamp ())));
433 fun add meths = NameSpace.extend_table (Sign.naming_of thy) (meths, new_meths)
434 handle Symtab.DUPS dups =>
435 error ("Duplicate declaration of method(s) " ^ commas_quote dups);
436 in MethodsData.map add thy end;
438 val add_method = add_methods o Library.single;
440 fun Method name meth cmt = Context.>> (add_method (name, meth, cmt));
445 fun method_setup (name, txt, cmt) =
447 "val thm = PureThy.get_thm_closure (Context.the_context ()) o PureThy.Name;\n\
448 \val thms = PureThy.get_thms_closure (Context.the_context ()) o PureThy.Name;\n\
449 \val method: bstring * (Method.src -> Proof.context -> Proof.method) * string"
450 "Method.add_method method"
451 ("(" ^ Library.quote name ^ ", " ^ txt ^ ", " ^ Library.quote cmt ^ ")");
455 (** concrete syntax **)
459 fun syntax scan = Args.context_syntax "method" scan;
461 fun simple_args scan f src ctxt : method =
462 #2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt);
464 fun ctxt_args (f: Proof.context -> method) src ctxt =
465 #2 (syntax (Scan.succeed (f ctxt)) src ctxt);
467 fun no_args m = ctxt_args (K m);
472 type modifier = (Proof.context -> Proof.context) * attribute;
476 fun sect ss = Scan.first (map Scan.lift ss);
477 fun thms ss = Scan.repeat (Scan.unless (sect ss) Attrib.multi_thm) >> flat;
479 fun app (f, att) (context, ths) = foldl_map att (Context.map_proof f context, ths);
481 fun section ss = (sect ss -- thms ss) :-- (fn (m, ths) => Scan.depend (fn context =>
482 Scan.succeed (app m (context, ths)))) >> #2;
484 fun sectioned args ss = args -- Scan.repeat (section ss);
488 fun sectioned_args args ss f src ctxt =
489 let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt
492 fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f;
493 fun bang_sectioned_args' ss scan f =
494 sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f);
495 fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f);
497 fun thms_ctxt_args f = sectioned_args (thms []) [] f;
498 fun thms_args f = thms_ctxt_args (K o f);
499 fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected");
508 val introN = "intro";
513 fun modifier name kind kind' att =
514 Args.$$$ name |-- (kind >> K NONE || kind' |-- Args.nat --| Args.colon >> SOME)
515 >> (pair (I: Proof.context -> Proof.context) o att);
517 val iprover_modifiers =
518 [modifier destN Args.bang_colon Args.bang ContextRules.dest_bang,
519 modifier destN Args.colon (Scan.succeed ()) ContextRules.dest,
520 modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang,
521 modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim,
522 modifier introN Args.bang_colon Args.bang ContextRules.intro_bang,
523 modifier introN Args.colon (Scan.succeed ()) ContextRules.intro,
524 Args.del -- Args.colon >> K (I, ContextRules.rule_del)];
529 bang_sectioned_args' iprover_modifiers (Scan.lift (Scan.option Args.nat))
530 (fn n => fn prems => fn ctxt => METHOD (fn facts =>
531 HEADGOAL (insert_tac (prems @ facts) THEN'
532 ObjectLogic.atomize_tac THEN' iprover_tac ctxt n)));
539 fun nat_thms_args f = uncurry f oo
540 (#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.thms));
542 fun goal_args' args tac src ctxt = #2 (syntax (Args.goal_spec HEADGOAL -- args >>
543 (fn (quant, s) => SIMPLE_METHOD'' quant (tac s))) src ctxt);
545 fun goal_args args tac = goal_args' (Scan.lift args) tac;
547 fun goal_args_ctxt' args tac src ctxt =
548 #2 (syntax (Args.goal_spec HEADGOAL -- args >>
549 (fn (quant, s) => SIMPLE_METHOD'' quant (tac ctxt s))) src ctxt);
551 fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac;
556 val _ = Context.add_setup (add_methods
557 [("fail", no_args fail, "force failure"),
558 ("succeed", no_args succeed, "succeed"),
559 ("-", no_args insert_facts, "do nothing (insert current facts only)"),
560 ("insert", thms_args insert, "insert theorems, ignoring facts (improper)"),
561 ("intro", thms_args intro, "repeatedly apply introduction rules"),
562 ("elim", thms_args elim, "repeatedly apply elimination rules"),
563 ("unfold", thms_ctxt_args unfold_meth, "unfold definitions"),
564 ("fold", thms_ctxt_args fold_meth, "fold definitions"),
565 ("atomize", (atomize o #2) oo syntax (Args.mode "full"),
566 "present local premises as object-level statements"),
567 ("iprover", iprover_meth, "intuitionistic proof search"),
568 ("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"),
569 ("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"),
570 ("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"),
571 ("frule", nat_thms_args frule, "apply rule in forward manner (improper)"),
572 ("this", no_args this, "apply current facts as rules"),
573 ("fact", thms_ctxt_args fact, "composition by facts from context"),
574 ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
575 ("rename_tac", goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac,
576 "rename parameters of goal (dynamic instantiation)"),
577 ("rotate_tac", goal_args (Scan.optional Args.int 1) Tactic.rotate_tac,
578 "rotate assumptions of goal"),
579 ("tactic", simple_args Args.name tactic, "ML tactic as proof method")]);
582 (*final declarations of this structure!*)
583 val unfold = unfold_meth;
584 val fold = fold_meth;
588 structure BasicMethod: BASIC_METHOD = Method;