src/Pure/Isar/method.ML
author wenzelm
Tue, 18 Dec 2007 19:54:33 +0100
changeset 25699 891fe6b71d3b
parent 25270 2ed7b34f58e6
child 25957 2cfb703fa8d8
permissions -rw-r--r--
named some critical sections;
     1 (*  Title:      Pure/Isar/method.ML
     2     ID:         $Id$
     3     Author:     Markus Wenzel, TU Muenchen
     4 
     5 Isar proof methods.
     6 *)
     7 
     8 signature BASIC_METHOD =
     9 sig
    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
    13   val trace_rules: bool ref
    14   val print_methods: theory -> unit
    15 end;
    16 
    17 signature METHOD =
    18 sig
    19   include BASIC_METHOD
    20   val apply: method -> thm list -> cases_tactic
    21   val RAW_METHOD_CASES: (thm list -> cases_tactic) -> method
    22   val RAW_METHOD: (thm list -> tactic) -> method
    23   val METHOD_CASES: (thm list -> cases_tactic) -> method
    24   val METHOD: (thm list -> tactic) -> method
    25   val fail: method
    26   val succeed: method
    27   val insert_tac: thm list -> int -> tactic
    28   val insert: thm list -> method
    29   val insert_facts: method
    30   val SIMPLE_METHOD: tactic -> method
    31   val SIMPLE_METHOD': (int -> 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 -> Proof.context -> method
    36   val intro: thm list -> method
    37   val elim: thm list -> method
    38   val unfold: thm list -> Proof.context -> method
    39   val fold: thm list -> Proof.context -> method
    40   val atomize: bool -> method
    41   val this: method
    42   val fact: thm list -> Proof.context -> method
    43   val assumption: Proof.context -> method
    44   val close: bool -> Proof.context -> method
    45   val trace: Proof.context -> thm list -> unit
    46   val rule_tac: thm list -> thm list -> int -> tactic
    47   val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic
    48   val intros_tac: thm list -> thm list -> 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
    56   type src
    57   datatype text =
    58     Basic of (Proof.context -> method) * Position.T |
    59     Source of src |
    60     Source_i of src |
    61     Then of text list |
    62     Orelse of text list |
    63     Try of text |
    64     Repeat1 of text |
    65     SelectGoals of int * text
    66   val primitive_text: (thm -> thm) -> text
    67   val succeed_text: text
    68   val default_text: text
    69   val this_text: text
    70   val done_text: text
    71   val sorry_text: bool -> text
    72   val finish_text: text option * bool -> Position.T -> text
    73   val method: theory -> src -> Proof.context -> method
    74   val method_i: theory -> src -> Proof.context -> method
    75   val add_methods: (bstring * (src -> Proof.context -> method) * string) list
    76     -> theory -> theory
    77   val add_method: bstring * (src -> Proof.context -> method) * string
    78     -> theory -> theory
    79   val method_setup: bstring * string * string -> theory -> theory
    80   val syntax: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list))
    81     -> src -> Proof.context -> 'a * Proof.context
    82   val simple_args: (Args.T list -> 'a * Args.T list)
    83     -> ('a -> Proof.context -> method) -> src -> Proof.context -> method
    84   val ctxt_args: (Proof.context -> method) -> src -> Proof.context -> method
    85   val no_args: method -> src -> Proof.context -> method
    86   type modifier
    87   val sectioned_args: (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
    88     (Args.T list -> modifier * Args.T list) list ->
    89     ('a -> Proof.context -> 'b) -> src -> Proof.context -> 'b
    90   val bang_sectioned_args:
    91     (Args.T list -> modifier * Args.T list) list ->
    92     (thm list -> Proof.context -> 'a) -> src -> Proof.context -> 'a
    93   val bang_sectioned_args':
    94     (Args.T list -> modifier * Args.T list) list ->
    95     (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
    96     ('a -> thm list -> Proof.context -> 'b) -> src -> Proof.context -> 'b
    97   val only_sectioned_args:
    98     (Args.T list -> modifier * Args.T list) list ->
    99     (Proof.context -> 'a) -> src -> Proof.context -> 'a
   100   val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> src ->
   101     Proof.context -> 'a
   102   val thms_args: (thm list -> 'a) -> src -> Proof.context -> 'a
   103   val thm_args: (thm -> 'a) -> src -> Proof.context -> 'a
   104   val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic)
   105     -> src -> Proof.context -> method
   106   val goal_args': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list))
   107     -> ('a -> int -> tactic) -> src -> Proof.context -> method
   108   val goal_args_ctxt: (Args.T list -> 'a * Args.T list) ->
   109     (Proof.context -> 'a -> int -> tactic) -> src -> Proof.context -> method
   110   val goal_args_ctxt': (Context.generic * Args.T list -> 'a * (Context.generic * Args.T list)) ->
   111     (Proof.context -> 'a -> int -> tactic) -> src -> Proof.context -> method
   112   val parse: OuterLex.token list -> text * OuterLex.token list
   113 end;
   114 
   115 structure Method: METHOD =
   116 struct
   117 
   118 (** generic tools **)
   119 
   120 (* goal addressing *)
   121 
   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;
   125 
   126 fun HEADGOAL tac = tac 1;
   127 
   128 
   129 
   130 (** proof methods **)
   131 
   132 (* datatype method *)
   133 
   134 datatype method = Meth of thm list -> cases_tactic;
   135 
   136 fun apply (Meth m) = m;
   137 
   138 val RAW_METHOD_CASES = Meth;
   139 
   140 fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac);
   141 
   142 fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts =>
   143   Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts));
   144 
   145 fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Goal.conjunction_tac THEN tac facts);
   146 
   147 val fail = METHOD (K no_tac);
   148 val succeed = METHOD (K all_tac);
   149 
   150 
   151 (* insert facts *)
   152 
   153 local
   154 
   155 fun cut_rule_tac rule =
   156   Tactic.rtac (Drule.forall_intr_vars rule COMP_INCR revcut_rl);
   157 
   158 in
   159 
   160 fun insert_tac [] i = all_tac
   161   | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts);
   162 
   163 val insert_facts = METHOD (ALLGOALS o insert_tac);
   164 fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms));
   165 
   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;
   169 
   170 end;
   171 
   172 
   173 (* shuffle subgoals *)
   174 
   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)));
   177 
   178 
   179 (* cheating *)
   180 
   181 fun cheating int ctxt = METHOD (K (setmp quick_and_dirty (int orelse ! quick_and_dirty)
   182     (SkipProof.cheat_tac (ProofContext.theory_of ctxt))));
   183 
   184 
   185 (* unfold intro/elim rules *)
   186 
   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));
   189 
   190 
   191 (* unfold/fold definitions *)
   192 
   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));
   195 
   196 
   197 (* atomize rule statements *)
   198 
   199 fun atomize false = SIMPLE_METHOD' (CHANGED_PROP o ObjectLogic.atomize_prems_tac)
   200   | atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac)));
   201 
   202 
   203 (* this -- resolve facts directly *)
   204 
   205 val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac));
   206 
   207 
   208 (* fact -- composition by facts from context *)
   209 
   210 fun fact [] ctxt = SIMPLE_METHOD' (ProofContext.some_fact_tac ctxt)
   211   | fact rules _ = SIMPLE_METHOD' (ProofContext.fact_tac rules);
   212 
   213 
   214 (* assumption *)
   215 
   216 local
   217 
   218 fun cond_rtac cond rule = SUBGOAL (fn (prop, i) =>
   219   if cond (Logic.strip_assums_concl prop)
   220   then Tactic.rtac rule i else no_tac);
   221 
   222 fun legacy_tac ctxt st =
   223   (legacy_feature ("implicit use of prems in assumption proof" ^ ContextPosition.str_of ctxt);
   224     all_tac st);
   225 
   226 fun assm_tac ctxt =
   227   assume_tac APPEND'
   228   Goal.assume_rule_tac ctxt APPEND'
   229   (CHANGED o solve_tac (Assumption.prems_of ctxt) THEN' K (legacy_tac ctxt)) APPEND'
   230   cond_rtac (can Logic.dest_equals) Drule.reflexive_thm APPEND'
   231   cond_rtac (can Logic.dest_term) Drule.termI;
   232 
   233 in
   234 
   235 fun assumption ctxt = METHOD (HEADGOAL o
   236   (fn [] => assm_tac ctxt
   237     | [fact] => solve_tac [fact]
   238     | _ => K no_tac));
   239 
   240 fun close immed ctxt = METHOD (K
   241   (FILTER Thm.no_prems
   242     ((if immed then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac)));
   243 
   244 end;
   245 
   246 
   247 (* rule etc. -- single-step refinements *)
   248 
   249 val trace_rules = ref false;
   250 
   251 fun trace ctxt rules =
   252   if ! trace_rules andalso not (null rules) then
   253     Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules)
   254     |> Pretty.string_of |> tracing
   255   else ();
   256 
   257 local
   258 
   259 fun gen_rule_tac tac rules facts =
   260   (fn i => fn st =>
   261     if null facts then tac rules i st
   262     else Seq.maps (fn rule => (tac o single) rule i st) (Drule.multi_resolves facts rules))
   263   THEN_ALL_NEW Goal.norm_hhf_tac;
   264 
   265 fun gen_arule_tac tac j rules facts =
   266   EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac);
   267 
   268 fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) =>
   269   let
   270     val rules =
   271       if not (null arg_rules) then arg_rules
   272       else flat (ContextRules.find_rules false facts goal ctxt)
   273   in trace ctxt rules; tac rules facts i end);
   274 
   275 fun meth tac x = METHOD (HEADGOAL o tac x);
   276 fun meth' tac x y = METHOD (HEADGOAL o tac x y);
   277 
   278 in
   279 
   280 val rule_tac = gen_rule_tac Tactic.resolve_tac;
   281 val rule = meth rule_tac;
   282 val some_rule_tac = gen_some_rule_tac rule_tac;
   283 val some_rule = meth' some_rule_tac;
   284 
   285 val erule = meth' (gen_arule_tac Tactic.eresolve_tac);
   286 val drule = meth' (gen_arule_tac Tactic.dresolve_tac);
   287 val frule = meth' (gen_arule_tac Tactic.forward_tac);
   288 
   289 end;
   290 
   291 
   292 (* intros_tac -- pervasive search spanned by intro rules *)
   293 
   294 fun intros_tac intros facts =
   295   ALLGOALS (insert_tac facts THEN'
   296       REPEAT_ALL_NEW (resolve_tac intros))
   297     THEN Tactic.distinct_subgoals_tac;
   298 
   299 
   300 (* iprover -- intuitionistic proof search *)
   301 
   302 local
   303 
   304 val remdups_tac = SUBGOAL (fn (g, i) =>
   305   let val prems = Logic.strip_assums_hyp g in
   306     REPEAT_DETERM_N (length prems - length (distinct op aconv prems))
   307     (Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i)
   308   end);
   309 
   310 fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac;
   311 
   312 val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist;
   313 
   314 fun safe_step_tac ctxt =
   315   ContextRules.Swrap ctxt
   316    (eq_assume_tac ORELSE'
   317     bires_tac true (ContextRules.netpair_bang ctxt));
   318 
   319 fun unsafe_step_tac ctxt =
   320   ContextRules.wrap ctxt
   321    (assume_tac APPEND'
   322     bires_tac false (ContextRules.netpair_bang ctxt) APPEND'
   323     bires_tac false (ContextRules.netpair ctxt));
   324 
   325 fun step_tac ctxt i =
   326   REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE
   327   REMDUPS (unsafe_step_tac ctxt) i;
   328 
   329 fun intprover_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else
   330   let
   331     val ps = Logic.strip_assums_hyp g;
   332     val c = Logic.strip_assums_concl g;
   333   in
   334     if member (fn ((ps1, c1), (ps2, c2)) =>
   335         c1 aconv c2 andalso
   336         length ps1 = length ps2 andalso
   337         gen_eq_set (op aconv) (ps1, ps2)) gs (ps, c) then no_tac
   338     else (step_tac ctxt THEN_ALL_NEW intprover_tac ctxt ((ps, c) :: gs) (d + 1) lim) i
   339   end);
   340 
   341 in
   342 
   343 fun iprover_tac ctxt opt_lim =
   344   SELECT_GOAL (DEEPEN (2, the_default 20 opt_lim) (intprover_tac ctxt [] 0) 4 1);
   345 
   346 end;
   347 
   348 
   349 (* ML tactics *)
   350 
   351 val tactic_ref = ref ((fn _ => raise Match): Proof.context -> thm list -> tactic);
   352 fun set_tactic f = tactic_ref := f;
   353 
   354 fun ml_tactic txt ctxt = NAMED_CRITICAL "ML" (fn () =>
   355   (ML_Context.use_mltext
   356     ("let fun tactic (ctxt: Proof.context) (facts: thm list) : tactic =\n"
   357       ^ txt ^ "\nin Method.set_tactic tactic end") false (SOME (Context.Proof ctxt));
   358     ML_Context.setmp (SOME (Context.Proof ctxt)) (! tactic_ref ctxt)));
   359 
   360 fun tactic txt ctxt = METHOD (ml_tactic txt ctxt);
   361 fun raw_tactic txt ctxt = RAW_METHOD (ml_tactic txt ctxt);
   362 
   363 
   364 
   365 (** method syntax **)
   366 
   367 (* method text *)
   368 
   369 type src = Args.src;
   370 
   371 datatype text =
   372   Basic of (Proof.context -> method) * Position.T |
   373   Source of src |
   374   Source_i of src |
   375   Then of text list |
   376   Orelse of text list |
   377   Try of text |
   378   Repeat1 of text |
   379   SelectGoals of int * text;
   380 
   381 fun primitive_text r = Basic (K (SIMPLE_METHOD (PRIMITIVE r)), Position.none);
   382 val succeed_text = Basic (K succeed, Position.none);
   383 val default_text = Source (Args.src (("default", []), Position.none));
   384 val this_text = Basic (K this, Position.none);
   385 val done_text = Basic (K (SIMPLE_METHOD all_tac), Position.none);
   386 fun sorry_text int = Basic (cheating int, Position.none);
   387 
   388 fun finish_text (NONE, immed) pos = Basic (close immed, pos)
   389   | finish_text (SOME txt, immed) pos = Then [txt, Basic (close immed, pos)];
   390 
   391 
   392 (* method definitions *)
   393 
   394 structure Methods = TheoryDataFun
   395 (
   396   type T = (((src -> Proof.context -> method) * string) * stamp) NameSpace.table;
   397   val empty = NameSpace.empty_table;
   398   val copy = I;
   399   val extend = I;
   400   fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>
   401     error ("Attempt to merge different versions of method " ^ quote dup);
   402 );
   403 
   404 fun print_methods thy =
   405   let
   406     val meths = Methods.get thy;
   407     fun prt_meth (name, ((_, comment), _)) = Pretty.block
   408       [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
   409   in
   410     [Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))]
   411     |> Pretty.chunks |> Pretty.writeln
   412   end;
   413 
   414 
   415 (* get methods *)
   416 
   417 fun method_i thy =
   418   let
   419     val meths = #2 (Methods.get thy);
   420     fun meth src =
   421       let val ((name, _), pos) = Args.dest_src src in
   422         (case Symtab.lookup meths name of
   423           NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
   424         | SOME ((mth, _), _) => mth src)
   425       end;
   426   in meth end;
   427 
   428 fun method thy = method_i thy o Args.map_name (NameSpace.intern (#1 (Methods.get thy)));
   429 
   430 
   431 (* add method *)
   432 
   433 fun add_methods raw_meths thy =
   434   let
   435     val new_meths = raw_meths |> map (fn (name, f, comment) =>
   436       (name, ((f, comment), stamp ())));
   437 
   438     fun add meths = NameSpace.extend_table (Sign.naming_of thy) new_meths meths
   439       handle Symtab.DUP dup => error ("Duplicate declaration of method " ^ quote dup);
   440   in Methods.map add thy end;
   441 
   442 val add_method = add_methods o Library.single;
   443 
   444 
   445 (* method_setup *)
   446 
   447 fun method_setup (name, txt, cmt) =
   448   ML_Context.use_let
   449     "val method: bstring * (Method.src -> Proof.context -> Proof.method) * string"
   450     "Context.map_theory (Method.add_method method)"
   451     ("(" ^ quote name ^ ", " ^ txt ^ ", " ^ quote cmt ^ ")")
   452   |> Context.theory_map;
   453 
   454 
   455 
   456 (** concrete syntax **)
   457 
   458 (* basic *)
   459 
   460 fun syntax scan = Args.context_syntax "method" scan;
   461 
   462 fun simple_args scan f src ctxt : method =
   463   fst (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt);
   464 
   465 fun ctxt_args (f: Proof.context -> method) src ctxt =
   466   fst (syntax (Scan.succeed (f ctxt)) src ctxt);
   467 
   468 fun no_args m = ctxt_args (K m);
   469 
   470 
   471 (* sections *)
   472 
   473 type modifier = (Proof.context -> Proof.context) * attribute;
   474 
   475 local
   476 
   477 fun thms ss = Scan.repeat (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm) >> flat;
   478 fun app (f, att) (context, ths) = foldl_map att (Context.map_proof f context, ths);
   479 
   480 fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|--
   481   (fn (m, ths) => Scan.succeed (app m (context, ths))));
   482 
   483 fun sectioned args ss = args -- Scan.repeat (section ss);
   484 
   485 in
   486 
   487 fun sectioned_args args ss f src ctxt =
   488   let val ((x, _), ctxt') = syntax (sectioned args ss) src ctxt
   489   in f x ctxt' end;
   490 
   491 fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f;
   492 fun bang_sectioned_args' ss scan f =
   493   sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f);
   494 fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f);
   495 
   496 fun thms_ctxt_args f = sectioned_args (thms []) [] f;
   497 fun thms_args f = thms_ctxt_args (K o f);
   498 fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected");
   499 
   500 end;
   501 
   502 
   503 (* iprover syntax *)
   504 
   505 local
   506 
   507 val introN = "intro";
   508 val elimN = "elim";
   509 val destN = "dest";
   510 val ruleN = "rule";
   511 
   512 fun modifier name kind kind' att =
   513   Args.$$$ name |-- (kind >> K NONE || kind' |-- Args.nat --| Args.colon >> SOME)
   514     >> (pair (I: Proof.context -> Proof.context) o att);
   515 
   516 val iprover_modifiers =
   517  [modifier destN Args.bang_colon Args.bang ContextRules.dest_bang,
   518   modifier destN Args.colon (Scan.succeed ()) ContextRules.dest,
   519   modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang,
   520   modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim,
   521   modifier introN Args.bang_colon Args.bang ContextRules.intro_bang,
   522   modifier introN Args.colon (Scan.succeed ()) ContextRules.intro,
   523   Args.del -- Args.colon >> K (I, ContextRules.rule_del)];
   524 
   525 in
   526 
   527 val iprover_meth =
   528   bang_sectioned_args' iprover_modifiers (Scan.lift (Scan.option Args.nat))
   529     (fn n => fn prems => fn ctxt => METHOD (fn facts =>
   530       HEADGOAL (insert_tac (prems @ facts) THEN'
   531       ObjectLogic.atomize_prems_tac THEN' iprover_tac ctxt n)));
   532 
   533 end;
   534 
   535 
   536 (* tactic syntax *)
   537 
   538 fun nat_thms_args f = uncurry f oo
   539   (fst oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.thms));
   540 
   541 fun goal_args' args tac src ctxt = fst (syntax (Args.goal_spec HEADGOAL -- args >>
   542   (fn (quant, s) => SIMPLE_METHOD'' quant (tac s))) src ctxt);
   543 
   544 fun goal_args args tac = goal_args' (Scan.lift args) tac;
   545 
   546 fun goal_args_ctxt' args tac src ctxt =
   547   fst (syntax (Args.goal_spec HEADGOAL -- args >>
   548   (fn (quant, s) => SIMPLE_METHOD'' quant (tac ctxt s))) src ctxt);
   549 
   550 fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac;
   551 
   552 
   553 (* theory setup *)
   554 
   555 val _ = Context.add_setup (add_methods
   556  [("fail", no_args fail, "force failure"),
   557   ("succeed", no_args succeed, "succeed"),
   558   ("-", no_args insert_facts, "do nothing (insert current facts only)"),
   559   ("insert", thms_args insert, "insert theorems, ignoring facts (improper)"),
   560   ("intro", thms_args intro, "repeatedly apply introduction rules"),
   561   ("elim", thms_args elim, "repeatedly apply elimination rules"),
   562   ("unfold", thms_ctxt_args unfold_meth, "unfold definitions"),
   563   ("fold", thms_ctxt_args fold_meth, "fold definitions"),
   564   ("atomize", (atomize o fst) oo syntax (Args.mode "full"),
   565     "present local premises as object-level statements"),
   566   ("iprover", iprover_meth, "intuitionistic proof search"),
   567   ("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"),
   568   ("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"),
   569   ("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"),
   570   ("frule", nat_thms_args frule, "apply rule in forward manner (improper)"),
   571   ("this", no_args this, "apply current facts as rules"),
   572   ("fact", thms_ctxt_args fact, "composition by facts from context"),
   573   ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
   574   ("rename_tac", goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac,
   575     "rename parameters of goal"),
   576   ("rotate_tac", goal_args (Scan.optional Args.int 1) Tactic.rotate_tac,
   577     "rotate assumptions of goal"),
   578   ("tactic", simple_args Args.name tactic, "ML tactic as proof method"),
   579   ("raw_tactic", simple_args Args.name raw_tactic, "ML tactic as raw proof method")]);
   580 
   581 
   582 (* outer parser *)
   583 
   584 local
   585 
   586 structure T = OuterLex;
   587 structure P = OuterParse;
   588 
   589 fun is_symid_meth s =
   590   s <> "|" andalso s <> "?" andalso s <> "+" andalso T.is_sid s;
   591 
   592 fun meth4 x =
   593  (P.position (P.xname >> rpair []) >> (Source o Args.src) ||
   594   P.$$$ "(" |-- P.!!! (meth0 --| P.$$$ ")")) x
   595 and meth3 x =
   596  (meth4 --| P.$$$ "?" >> Try ||
   597   meth4 --| P.$$$ "+" >> Repeat1 ||
   598   meth4 -- (P.$$$ "[" |-- Scan.optional P.nat 1 --| P.$$$ "]") >> (SelectGoals o swap) ||
   599   meth4) x
   600 and meth2 x =
   601  (P.position (P.xname -- P.args1 is_symid_meth false) >> (Source o Args.src) ||
   602   meth3) x
   603 and meth1 x = (P.enum1 "," meth2 >> (fn [m] => m | ms => Then ms)) x
   604 and meth0 x = (P.enum1 "|" meth1 >> (fn [m] => m | ms => Orelse ms)) x;
   605 
   606 in val parse = meth3 end;
   607 
   608 
   609 (*final declarations of this structure!*)
   610 val unfold = unfold_meth;
   611 val fold = fold_meth;
   612 
   613 end;
   614 
   615 structure BasicMethod: BASIC_METHOD = Method;
   616 open BasicMethod;