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