src/Pure/Isar/method.ML
author ballarin
Fri, 02 Apr 2004 14:08:30 +0200
changeset 14508 859b11514537
parent 14215 ebf291f3b449
child 14718 f52f2cf2d137
permissions -rw-r--r--
Experimental command for instantiation of locales in proof contexts:
instantiate <label>: <loc>
     1 (*  Title:      Pure/Isar/method.ML
     2     ID:         $Id$
     3     Author:     Markus Wenzel, TU Muenchen
     4     License:    GPL (GNU GENERAL PUBLIC LICENSE)
     5 
     6 Proof methods.
     7 *)
     8 
     9 signature BASIC_METHOD =
    10 sig
    11   val trace_rules: bool ref
    12   val print_methods: theory -> unit
    13   val Method: bstring -> (Args.src -> Proof.context -> Proof.method) -> string -> unit
    14 end;
    15 
    16 signature METHOD =
    17 sig
    18   include BASIC_METHOD
    19   val trace: Proof.context -> thm list -> unit
    20   val RAW_METHOD: (thm list -> tactic) -> Proof.method
    21   val RAW_METHOD_CASES:
    22     (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method
    23   val METHOD: (thm list -> tactic) -> Proof.method
    24   val METHOD_CASES:
    25     (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> Proof.method
    26   val SIMPLE_METHOD: tactic -> Proof.method
    27   val SIMPLE_METHOD': ((int -> tactic) -> tactic) -> (int -> tactic) -> Proof.method
    28   val fail: Proof.method
    29   val succeed: Proof.method
    30   val defer: int option -> Proof.method
    31   val prefer: int -> Proof.method
    32   val insert_tac: thm list -> int -> tactic
    33   val insert: thm list -> Proof.method
    34   val insert_facts: Proof.method
    35   val unfold: thm list -> Proof.method
    36   val fold: thm list -> Proof.method
    37   val multi_resolve: thm list -> thm -> thm Seq.seq
    38   val multi_resolves: thm list -> thm list -> thm Seq.seq
    39   val rules_tac: Proof.context -> int option -> int -> tactic
    40   val rule_tac: thm list -> thm list -> int -> tactic
    41   val some_rule_tac: thm list -> Proof.context -> thm list -> int -> tactic
    42   val rule: thm list -> Proof.method
    43   val erule: int -> thm list -> Proof.method
    44   val drule: int -> thm list -> Proof.method
    45   val frule: int -> thm list -> Proof.method
    46   val this: Proof.method
    47   val assumption: Proof.context -> Proof.method
    48   val bires_inst_tac: bool -> Proof.context -> (indexname * string) list -> thm -> int -> tactic
    49   val set_tactic: (Proof.context -> thm list -> tactic) -> unit
    50   val tactic: string -> Proof.context -> Proof.method
    51   exception METHOD_FAIL of (string * Position.T) * exn
    52   val method: theory -> Args.src -> Proof.context -> Proof.method
    53   val add_method: bstring * (Args.src -> Proof.context -> Proof.method) * string
    54     -> theory -> theory
    55   val add_methods: (bstring * (Args.src -> Proof.context -> Proof.method) * string) list
    56     -> theory -> theory
    57   val syntax: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ->
    58     Args.src -> Proof.context -> Proof.context * 'a
    59   val simple_args: (Args.T list -> 'a * Args.T list)
    60     -> ('a -> Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
    61   val ctxt_args: (Proof.context -> Proof.method) -> Args.src -> Proof.context -> Proof.method
    62   val no_args: Proof.method -> Args.src -> Proof.context -> Proof.method
    63   type modifier
    64   val sectioned_args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ->
    65     (Args.T list -> modifier * Args.T list) list ->
    66     ('a -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b
    67   val bang_sectioned_args:
    68     (Args.T list -> modifier * Args.T list) list ->
    69     (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a
    70   val bang_sectioned_args':
    71     (Args.T list -> modifier * Args.T list) list ->
    72     (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) ->
    73     ('a -> thm list -> Proof.context -> 'b) -> Args.src -> Proof.context -> 'b
    74   val only_sectioned_args:
    75     (Args.T list -> modifier * Args.T list) list ->
    76     (Proof.context -> 'a) -> Args.src -> Proof.context -> 'a
    77   val thms_ctxt_args: (thm list -> Proof.context -> 'a) -> Args.src -> Proof.context -> 'a
    78   val thms_args: (thm list -> 'a) -> Args.src -> Proof.context -> 'a
    79   val thm_args: (thm -> 'a) -> Args.src -> Proof.context -> 'a
    80   datatype text =
    81     Basic of (Proof.context -> Proof.method) |
    82     Source of Args.src |
    83     Then of text list |
    84     Orelse of text list |
    85     Try of text |
    86     Repeat1 of text
    87   val refine: text -> Proof.state -> Proof.state Seq.seq
    88   val refine_end: text -> Proof.state -> Proof.state Seq.seq
    89   val proof: text option -> Proof.state -> Proof.state Seq.seq
    90   val local_qed: bool -> text option
    91     -> (Proof.context -> string * (string * thm list) list -> unit) *
    92       (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq
    93   val local_terminal_proof: text * text option
    94     -> (Proof.context -> string * (string * thm list) list -> unit) *
    95       (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq
    96   val local_default_proof: (Proof.context -> string * (string * thm list) list -> unit) *
    97     (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq
    98   val local_immediate_proof: (Proof.context -> string * (string * thm list) list -> unit) *
    99     (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq
   100   val local_done_proof: (Proof.context -> string * (string * thm list) list -> unit) *
   101     (Proof.context -> thm -> unit) -> Proof.state -> Proof.state Seq.seq
   102   val global_qed: bool -> text option
   103     -> Proof.state -> theory * ((string * string) * (string * thm list) list)
   104   val global_terminal_proof: text * text option
   105     -> Proof.state -> theory * ((string * string) * (string * thm list) list)
   106   val global_default_proof: Proof.state -> theory * ((string * string) * (string * thm list) list)
   107   val global_immediate_proof: Proof.state ->
   108     theory * ((string * string) * (string * thm list) list)
   109   val global_done_proof: Proof.state -> theory * ((string * string) * (string * thm list) list)
   110   val goal_args: (Args.T list -> 'a * Args.T list) -> ('a -> int -> tactic)
   111     -> Args.src -> Proof.context -> Proof.method
   112   val goal_args': (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))
   113     -> ('a -> int -> tactic) -> Args.src -> Proof.context -> Proof.method
   114   val goal_args_ctxt: (Args.T list -> 'a * Args.T list) -> (Proof.context -> 'a -> int -> tactic)
   115     -> Args.src -> Proof.context -> Proof.method
   116   val goal_args_ctxt': (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))
   117     -> (Proof.context -> 'a -> int -> tactic) -> Args.src -> Proof.context -> Proof.method
   118   val setup: (theory -> theory) list
   119 end;
   120 
   121 structure Method: METHOD =
   122 struct
   123 
   124 
   125 (** proof methods **)
   126 
   127 (* tracing *)
   128 
   129 val trace_rules = ref false;
   130 
   131 fun trace ctxt rules =
   132   conditional (! trace_rules andalso not (null rules)) (fn () =>
   133     Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules)
   134     |> Pretty.string_of |> tracing);
   135 
   136 
   137 (* make methods *)
   138 
   139 val RAW_METHOD = Proof.method;
   140 val RAW_METHOD_CASES = Proof.method_cases;
   141 
   142 fun METHOD m = Proof.method (fn facts => TRY Tactic.conjunction_tac THEN m facts);
   143 fun METHOD_CASES m =
   144   Proof.method_cases (fn facts => Seq.THEN (TRY Tactic.conjunction_tac, m facts));
   145 
   146 
   147 (* primitive *)
   148 
   149 val fail = METHOD (K no_tac);
   150 val succeed = METHOD (K all_tac);
   151 
   152 
   153 (* shuffle *)
   154 
   155 fun prefer i = METHOD (K (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1)));
   156 fun defer opt_i = METHOD (K (Tactic.defer_tac (if_none opt_i 1)));
   157 
   158 
   159 (* insert *)
   160 
   161 local
   162 
   163 fun cut_rule_tac raw_rule =
   164   let
   165     val rule = Drule.forall_intr_vars raw_rule;
   166     val revcut_rl = Drule.incr_indexes_wrt [] [] [] [rule] Drule.revcut_rl;
   167   in Tactic.rtac (rule COMP revcut_rl) end;
   168 
   169 in
   170 
   171 fun insert_tac [] i = all_tac
   172   | insert_tac facts i = EVERY (map (fn th => cut_rule_tac th i) facts);
   173 
   174 val insert_facts = METHOD (ALLGOALS o insert_tac);
   175 fun insert thms = METHOD (fn _ => ALLGOALS (insert_tac thms));
   176 
   177 fun SIMPLE_METHOD tac = METHOD (fn facts => ALLGOALS (insert_tac facts) THEN tac);
   178 fun SIMPLE_METHOD' quant tac = METHOD (fn facts => quant (insert_tac facts THEN' tac));
   179 
   180 end;
   181 
   182 
   183 (* unfold/fold definitions *)
   184 
   185 fun unfold ths = SIMPLE_METHOD (CHANGED_PROP (rewrite_goals_tac ths));
   186 fun fold ths = SIMPLE_METHOD (CHANGED_PROP (fold_goals_tac ths));
   187 
   188 
   189 (* atomize rule statements *)
   190 
   191 fun atomize false = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o ObjectLogic.atomize_tac)
   192   | atomize true = RAW_METHOD (K (HEADGOAL (CHANGED_PROP o ObjectLogic.full_atomize_tac)));
   193 
   194 
   195 (* unfold intro/elim rules *)
   196 
   197 fun intro ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.match_tac ths));
   198 fun elim ths = SIMPLE_METHOD' HEADGOAL (CHANGED_PROP o REPEAT_ALL_NEW (Tactic.ematch_tac ths));
   199 
   200 
   201 (* multi_resolve *)
   202 
   203 local
   204 
   205 fun res th i rule =
   206   Thm.biresolution false [(false, th)] i rule handle THM _ => Seq.empty;
   207 
   208 fun multi_res _ [] rule = Seq.single rule
   209   | multi_res i (th :: ths) rule = Seq.flat (Seq.map (res th i) (multi_res (i + 1) ths rule));
   210 
   211 in
   212 
   213 val multi_resolve = multi_res 1;
   214 fun multi_resolves facts rules = Seq.flat (Seq.map (multi_resolve facts) (Seq.of_list rules));
   215 
   216 end;
   217 
   218 
   219 (* rules_tac *)
   220 
   221 local
   222 
   223 val remdups_tac = SUBGOAL (fn (g, i) =>
   224   let val prems = Logic.strip_assums_hyp g in
   225     REPEAT_DETERM_N (length prems - length (gen_distinct op aconv prems))
   226     (Tactic.ematch_tac [Drule.remdups_rl] i THEN Tactic.eq_assume_tac i)
   227   end);
   228 
   229 fun REMDUPS tac = tac THEN_ALL_NEW remdups_tac;
   230 
   231 fun gen_eq_set e s1 s2 =
   232   length s1 = length s2 andalso
   233   gen_subset e (s1, s2) andalso gen_subset e (s2, s1);
   234 
   235 val bires_tac = Tactic.biresolution_from_nets_tac ContextRules.orderlist;
   236 
   237 fun safe_step_tac ctxt =
   238   ContextRules.Swrap ctxt
   239    (eq_assume_tac ORELSE'
   240     bires_tac true (ContextRules.netpair_bang ctxt));
   241 
   242 fun unsafe_step_tac ctxt =
   243   ContextRules.wrap ctxt
   244    (assume_tac APPEND'
   245     bires_tac false (ContextRules.netpair_bang ctxt) APPEND'
   246     bires_tac false (ContextRules.netpair ctxt));
   247 
   248 fun step_tac ctxt i =
   249   REPEAT_DETERM1 (REMDUPS (safe_step_tac ctxt) i) ORELSE
   250   REMDUPS (unsafe_step_tac ctxt) i;
   251 
   252 fun intpr_tac ctxt gs d lim = SUBGOAL (fn (g, i) => if d > lim then no_tac else
   253   let
   254     val ps = Logic.strip_assums_hyp g;
   255     val c = Logic.strip_assums_concl g;
   256   in
   257     if gen_mem (fn ((ps1, c1), (ps2, c2)) =>
   258       c1 aconv c2 andalso gen_eq_set op aconv ps1 ps2) ((ps, c), gs) then no_tac
   259     else (step_tac ctxt THEN_ALL_NEW intpr_tac ctxt ((ps, c) :: gs) (d + 1) lim) i
   260   end);
   261 
   262 in
   263 
   264 fun rules_tac ctxt opt_lim =
   265   SELECT_GOAL (DEEPEN (2, if_none opt_lim 20) (intpr_tac ctxt [] 0) 4 1);
   266 
   267 end;
   268 
   269 
   270 (* rule_tac etc. *)
   271 
   272 local
   273 
   274 fun gen_rule_tac tac rules [] i st = tac rules i st
   275   | gen_rule_tac tac rules facts i st =
   276       Seq.flat (Seq.map (fn rule => (tac o single) rule i st) (multi_resolves facts rules));
   277 
   278 fun gen_arule_tac tac j rules facts =
   279   EVERY' (gen_rule_tac tac rules facts :: replicate j Tactic.assume_tac);
   280 
   281 fun gen_some_rule_tac tac arg_rules ctxt facts = SUBGOAL (fn (goal, i) =>
   282   let
   283     val rules =
   284       if not (null arg_rules) then arg_rules
   285       else flat (ContextRules.find_rules false facts goal ctxt)
   286   in trace ctxt rules; tac rules facts i end);
   287 
   288 fun meth tac x = METHOD (HEADGOAL o tac x);
   289 fun meth' tac x y = METHOD (HEADGOAL o tac x y);
   290 
   291 in
   292 
   293 val rule_tac = gen_rule_tac Tactic.resolve_tac;
   294 val rule = meth rule_tac;
   295 val some_rule_tac = gen_some_rule_tac rule_tac;
   296 val some_rule = meth' some_rule_tac;
   297 
   298 val erule = meth' (gen_arule_tac Tactic.eresolve_tac);
   299 val drule = meth' (gen_arule_tac Tactic.dresolve_tac);
   300 val frule = meth' (gen_arule_tac Tactic.forward_tac);
   301 
   302 end;
   303 
   304 
   305 (* this *)
   306 
   307 val this = METHOD (EVERY o map (HEADGOAL o Tactic.rtac));
   308 
   309 
   310 (* assumption *)
   311 
   312 fun asm_tac ths =
   313   foldr (op APPEND') (map (fn th => Tactic.rtac th THEN_ALL_NEW assume_tac) ths, K no_tac);
   314 
   315 fun assm_tac ctxt =
   316   assume_tac APPEND'
   317   asm_tac (ProofContext.prems_of ctxt) APPEND'
   318   Tactic.rtac Drule.reflexive_thm;
   319 
   320 fun assumption_tac ctxt [] = assm_tac ctxt
   321   | assumption_tac _ [fact] = asm_tac [fact]
   322   | assumption_tac _ _ = K no_tac;
   323 
   324 fun assumption ctxt = METHOD (HEADGOAL o assumption_tac ctxt);
   325 
   326 
   327 (* res_inst_tac etc. *)
   328 
   329 (* Reimplemented to support both static (Isar) and dynamic (proof state)
   330    context.  By Clemens Ballarin. *)
   331 
   332 fun bires_inst_tac bires_flag ctxt insts thm =
   333   let
   334     val sign = ProofContext.sign_of ctxt;
   335     (* Separate type and term insts *)
   336     fun has_type_var ((x, _), _) = (case Symbol.explode x of
   337           "'"::cs => true | cs => false);
   338     val Tinsts = filter has_type_var insts;
   339     val tinsts = filter_out has_type_var insts;
   340     (* Tactic *)
   341     fun tac i st =
   342       let
   343         (* Preprocess state: extract environment information:
   344            - variables and their types
   345            - type variables and their sorts
   346            - parameters and their types *)
   347         val (types, sorts) = types_sorts st;
   348     (* Process type insts: Tinsts_env *)
   349     fun absent xi = error
   350 	  ("No such variable in theorem: " ^ Syntax.string_of_vname xi);
   351     val (rtypes, rsorts) = types_sorts thm;
   352     fun readT (xi, s) =
   353         let val S = case rsorts xi of Some S => S | None => absent xi;
   354             val T = Sign.read_typ (sign, sorts) s;
   355         in if Sign.typ_instance sign (T, TVar (xi, S)) then (xi, T)
   356            else error
   357              ("Instantiation of " ^ Syntax.string_of_vname xi ^ " fails")
   358         end;
   359     val Tinsts_env = map readT Tinsts;
   360     (* Preprocess rule: extract vars and their types, apply Tinsts *)
   361     fun get_typ xi =
   362       (case rtypes xi of
   363            Some T => typ_subst_TVars Tinsts_env T
   364          | None => absent xi);
   365     val (xis, ss) = Library.split_list tinsts;
   366     val Ts = map get_typ xis;
   367 	val (_, _, Bi, _) = dest_state(st,i)
   368 	val params = Logic.strip_params Bi
   369 			     (* params of subgoal i as string typ pairs *)
   370 	val params = rev(Term.rename_wrt_term Bi params)
   371 			   (* as they are printed: bound variables with *)
   372                            (* the same name are renamed during printing *)
   373         fun types' (a, ~1) = (case assoc (params, a) of
   374                 None => types (a, ~1)
   375               | some => some)
   376           | types' xi = types xi;
   377         val used = add_term_tvarnames
   378                   (prop_of st $ prop_of thm,[])
   379 	val (ts, envT) =
   380 	  ProofContext.read_termTs_env (types', sorts, used) ctxt (ss ~~ Ts);
   381 	val cenvT = map (apsnd (Thm.ctyp_of sign)) (envT @ Tinsts_env);
   382 	val cenv =
   383 	  map
   384 	    (fn (xi, t) =>
   385 	      pairself (Thm.cterm_of sign) (Var (xi, fastype_of t), t))
   386 	    (gen_distinct
   387 	      (fn ((x1, t1), (x2, t2)) => x1 = x2 andalso t1 aconv t2)
   388 	      (xis ~~ ts));
   389 	(* Lift and instantiate rule *)
   390 	val {maxidx, ...} = rep_thm st;
   391 	val paramTs = map #2 params
   392 	and inc = maxidx+1
   393 	fun liftvar (Var ((a,j), T)) =
   394 	      Var((a, j+inc), paramTs ---> incr_tvar inc T)
   395 	  | liftvar t = raise TERM("Variable expected", [t]);
   396 	fun liftterm t = list_abs_free
   397 	      (params, Logic.incr_indexes(paramTs,inc) t)
   398 	fun liftpair (cv,ct) =
   399 	      (cterm_fun liftvar cv, cterm_fun liftterm ct)
   400 	fun lifttvar((a,i),ctyp) =
   401 	    let val {T,sign} = rep_ctyp ctyp
   402 	    in  ((a,i+inc), ctyp_of sign (incr_tvar inc T)) end
   403 	val rule = Drule.instantiate
   404 	      (map lifttvar cenvT, map liftpair cenv)
   405 	      (lift_rule (st, i) thm)
   406       in
   407 	if i > nprems_of st then no_tac st
   408 	else st |>
   409 	  compose_tac (bires_flag, rule, nprems_of thm) i
   410       end
   411 	   handle TERM (msg,_)   => (warning msg; no_tac st)
   412 		| THM  (msg,_,_) => (warning msg; no_tac st);
   413   in
   414     tac
   415   end;
   416 
   417 fun gen_inst _ tac _ (quant, ([], thms)) =
   418       METHOD (fn facts => quant (insert_tac facts THEN' tac thms))
   419   | gen_inst inst_tac _ ctxt (quant, (insts, [thm])) =
   420       METHOD (fn facts =>
   421         quant (insert_tac facts THEN' inst_tac ctxt insts thm))
   422   | gen_inst _ _ _ _ = error "Cannot have instantiations with multiple rules";
   423       
   424 val res_inst_meth = gen_inst (bires_inst_tac false) Tactic.resolve_tac;
   425 
   426 val eres_inst_meth = gen_inst (bires_inst_tac true) Tactic.eresolve_tac;
   427 
   428 (* Preserve Var indexes of rl; increment revcut_rl instead.
   429    Copied from tactic.ML *)
   430 fun make_elim_preserve rl =
   431   let val {maxidx,...} = rep_thm rl
   432       fun cvar xi = cterm_of (Theory.sign_of ProtoPure.thy) (Var(xi,propT));
   433       val revcut_rl' =
   434           instantiate ([],  [(cvar("V",0), cvar("V",maxidx+1)),
   435                              (cvar("W",0), cvar("W",maxidx+1))]) revcut_rl
   436       val arg = (false, rl, nprems_of rl)
   437       val [th] = Seq.list_of (bicompose false arg 1 revcut_rl')
   438   in  th  end
   439   handle Bind => raise THM("make_elim_preserve", 1, [rl]);
   440 
   441 val cut_inst_meth =
   442   gen_inst
   443     (fn ctxt => fn insts => fn thm =>
   444        bires_inst_tac false ctxt insts (make_elim_preserve thm))
   445     Tactic.cut_rules_tac;
   446 
   447 val dres_inst_meth =
   448   gen_inst
   449     (fn ctxt => fn insts => fn rule =>
   450        bires_inst_tac true ctxt insts (make_elim_preserve rule))
   451     Tactic.dresolve_tac;
   452 
   453 val forw_inst_meth =
   454   gen_inst
   455     (fn ctxt => fn insts => fn rule =>
   456        bires_inst_tac false ctxt insts (make_elim_preserve rule) THEN'
   457        assume_tac)
   458     Tactic.forward_tac;
   459 
   460 fun subgoal_tac ctxt sprop =
   461   DETERM o bires_inst_tac false ctxt [(("psi", 0), sprop)] cut_rl THEN'
   462   SUBGOAL (fn (prop, _) =>
   463     let val concl' = Logic.strip_assums_concl prop in
   464       if null (term_tvars concl') then ()
   465       else warning "Type variables in new subgoal: add a type constraint?";
   466       all_tac
   467   end);
   468 
   469 fun subgoals_tac ctxt sprops = EVERY' (map (subgoal_tac ctxt) sprops);
   470 
   471 fun thin_tac ctxt s =
   472   bires_inst_tac true ctxt [(("V", 0), s)] thin_rl;
   473 
   474 (* simple Prolog interpreter *)
   475 
   476 fun prolog_tac rules facts =
   477   DEPTH_SOLVE_1 (HEADGOAL (Tactic.assume_tac APPEND' Tactic.resolve_tac (facts @ rules)));
   478 
   479 val prolog = METHOD o prolog_tac;
   480 
   481 
   482 (* ML tactics *)
   483 
   484 val tactic_ref = ref ((fn _ => raise Match): Proof.context -> thm list -> tactic);
   485 fun set_tactic f = tactic_ref := f;
   486 
   487 fun tactic txt ctxt = METHOD (fn facts =>
   488   (Context.use_mltext
   489     ("let fun tactic (ctxt: PureIsar.Proof.context) (facts: thm list) : tactic = \
   490      \let val thm = PureIsar.ProofContext.get_thm_closure ctxt\n\
   491      \  and thms = PureIsar.ProofContext.get_thms_closure ctxt in\n"
   492      ^ txt ^
   493      "\nend in PureIsar.Method.set_tactic tactic end")
   494     false None;
   495     Context.setmp (Some (ProofContext.theory_of ctxt)) (! tactic_ref ctxt) facts));
   496 
   497 
   498 
   499 (** methods theory data **)
   500 
   501 (* data kind 'Isar/methods' *)
   502 
   503 structure MethodsDataArgs =
   504 struct
   505   val name = "Isar/methods";
   506   type T =
   507     {space: NameSpace.T,
   508      meths: (((Args.src -> Proof.context -> Proof.method) * string) * stamp) Symtab.table};
   509 
   510   val empty = {space = NameSpace.empty, meths = Symtab.empty};
   511   val copy = I;
   512   val prep_ext = I;
   513   fun merge ({space = space1, meths = meths1}, {space = space2, meths = meths2}) =
   514     {space = NameSpace.merge (space1, space2),
   515       meths = Symtab.merge eq_snd (meths1, meths2) handle Symtab.DUPS dups =>
   516         error ("Attempt to merge different versions of methods " ^ commas_quote dups)};
   517 
   518   fun print _ {space, meths} =
   519     let
   520       fun prt_meth (name, ((_, comment), _)) = Pretty.block
   521         [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
   522     in
   523       [Pretty.big_list "methods:" (map prt_meth (NameSpace.cond_extern_table space meths))]
   524       |> Pretty.chunks |> Pretty.writeln
   525     end;
   526 end;
   527 
   528 structure MethodsData = TheoryDataFun(MethodsDataArgs);
   529 val print_methods = MethodsData.print;
   530 
   531 
   532 (* get methods *)
   533 
   534 exception METHOD_FAIL of (string * Position.T) * exn;
   535 
   536 fun method thy =
   537   let
   538     val {space, meths} = MethodsData.get thy;
   539 
   540     fun meth src =
   541       let
   542         val ((raw_name, _), pos) = Args.dest_src src;
   543         val name = NameSpace.intern space raw_name;
   544       in
   545         (case Symtab.lookup (meths, name) of
   546           None => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
   547         | Some ((mth, _), _) => transform_failure (curry METHOD_FAIL (name, pos)) (mth src))
   548       end;
   549   in meth end;
   550 
   551 
   552 (* add_method(s) *)
   553 
   554 fun add_methods raw_meths thy =
   555   let
   556     val full = Sign.full_name (Theory.sign_of thy);
   557     val new_meths =
   558       map (fn (name, f, comment) => (full name, ((f, comment), stamp ()))) raw_meths;
   559 
   560     val {space, meths} = MethodsData.get thy;
   561     val space' = NameSpace.extend (space, map fst new_meths);
   562     val meths' = Symtab.extend (meths, new_meths) handle Symtab.DUPS dups =>
   563       error ("Duplicate declaration of method(s) " ^ commas_quote dups);
   564   in
   565     thy |> MethodsData.put {space = space', meths = meths'}
   566   end;
   567 
   568 val add_method = add_methods o Library.single;
   569 
   570 (*implicit version*)
   571 fun Method name meth cmt = Context.>> (add_methods [(name, meth, cmt)]);
   572 
   573 
   574 
   575 (** method syntax **)
   576 
   577 (* basic *)
   578 
   579 fun syntax (scan: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list))) =
   580   Args.syntax "method" scan;
   581 
   582 fun simple_args scan f src ctxt : Proof.method =
   583   #2 (syntax (Scan.lift (scan >> (fn x => f x ctxt))) src ctxt);
   584 
   585 fun ctxt_args (f: Proof.context -> Proof.method) src ctxt =
   586   #2 (syntax (Scan.succeed (f ctxt)) src ctxt);
   587 
   588 fun no_args m = ctxt_args (K m);
   589 
   590 
   591 (* sections *)
   592 
   593 type modifier = (Proof.context -> Proof.context) * Proof.context attribute;
   594 
   595 local
   596 
   597 fun sect ss = Scan.first (map Scan.lift ss);
   598 fun thms ss = Scan.unless (sect ss) Attrib.local_thms;
   599 fun thmss ss = Scan.repeat (thms ss) >> flat;
   600 
   601 fun apply (f, att) (ctxt, ths) = Thm.applys_attributes ((f ctxt, ths), [att]);
   602 
   603 fun section ss = (sect ss -- thmss ss) :-- (fn (m, ths) => Scan.depend (fn ctxt =>
   604   Scan.succeed (apply m (ctxt, ths)))) >> #2;
   605 
   606 fun sectioned args ss = args -- Scan.repeat (section ss);
   607 
   608 in
   609 
   610 fun sectioned_args args ss f src ctxt =
   611   let val (ctxt', (x, _)) = syntax (sectioned args ss) src ctxt
   612   in f x ctxt' end;
   613 
   614 fun bang_sectioned_args ss f = sectioned_args Args.bang_facts ss f;
   615 fun bang_sectioned_args' ss scan f =
   616   sectioned_args (Args.bang_facts -- scan >> swap) ss (uncurry f);
   617 fun only_sectioned_args ss f = sectioned_args (Scan.succeed ()) ss (fn () => f);
   618 
   619 fun thms_ctxt_args f = sectioned_args (thmss []) [] f;
   620 fun thms_args f = thms_ctxt_args (K o f);
   621 fun thm_args f = thms_args (fn [thm] => f thm | _ => error "Single theorem expected");
   622 
   623 end;
   624 
   625 
   626 (* rules syntax *)
   627 
   628 local
   629 
   630 val introN = "intro";
   631 val elimN = "elim";
   632 val destN = "dest";
   633 val ruleN = "rule";
   634 
   635 fun modifier name kind kind' att =
   636   Args.$$$ name |-- (kind >> K None || kind' |-- Args.nat --| Args.colon >> Some)
   637     >> (pair (I: Proof.context -> Proof.context) o att);
   638 
   639 val rules_modifiers =
   640  [modifier destN Args.bang_colon Args.bang ContextRules.dest_bang_local,
   641   modifier destN Args.colon (Scan.succeed ()) ContextRules.dest_local,
   642   modifier elimN Args.bang_colon Args.bang ContextRules.elim_bang_local,
   643   modifier elimN Args.colon (Scan.succeed ()) ContextRules.elim_local,
   644   modifier introN Args.bang_colon Args.bang ContextRules.intro_bang_local,
   645   modifier introN Args.colon (Scan.succeed ()) ContextRules.intro_local,
   646   Args.del -- Args.colon >> K (I, ContextRules.rule_del_local)];
   647 
   648 in
   649 
   650 fun rules_args m = bang_sectioned_args' rules_modifiers (Scan.lift (Scan.option Args.nat)) m;
   651 
   652 fun rules_meth n prems ctxt = METHOD (fn facts =>
   653   HEADGOAL (insert_tac (prems @ facts) THEN' ObjectLogic.atomize_tac THEN' rules_tac ctxt n));
   654 
   655 end;
   656 
   657 
   658 (* tactic syntax *)
   659 
   660 fun nat_thms_args f = uncurry f oo
   661   (#2 oo syntax (Scan.lift (Scan.optional (Args.parens Args.nat) 0) -- Attrib.local_thmss));
   662 
   663 val insts =
   664   Scan.optional
   665     (Args.enum1 "and" (Scan.lift (Args.name -- (Args.$$$ "=" |-- Args.!!! Args.name))) --|
   666       Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss;
   667 
   668 fun inst_args f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts) src ctxt));
   669 
   670 val insts_var =
   671   Scan.optional
   672     (Args.enum1 "and" (Scan.lift (Args.var -- (Args.$$$ "=" |-- Args.!!! Args.name))) --|
   673       Scan.lift (Args.$$$ "in")) [] -- Attrib.local_thmss;
   674 
   675 fun inst_args_var f src ctxt = f ctxt (#2 (syntax (Args.goal_spec HEADGOAL -- insts_var) src ctxt));
   676 
   677 fun goal_args' args tac src ctxt = #2 (syntax (Args.goal_spec HEADGOAL -- args >>
   678   (fn (quant, s) => SIMPLE_METHOD' quant (tac s))) src ctxt);
   679 
   680 fun goal_args args tac = goal_args' (Scan.lift args) tac;
   681 
   682 fun goal_args_ctxt' args tac src ctxt =
   683   #2 (syntax (Args.goal_spec HEADGOAL -- args >>
   684   (fn (quant, s) => SIMPLE_METHOD' quant (tac ctxt s))) src ctxt);
   685 
   686 fun goal_args_ctxt args tac = goal_args_ctxt' (Scan.lift args) tac;
   687 
   688 (** method text **)
   689 
   690 (* datatype text *)
   691 
   692 datatype text =
   693   Basic of (Proof.context -> Proof.method) |
   694   Source of Args.src |
   695   Then of text list |
   696   Orelse of text list |
   697   Try of text |
   698   Repeat1 of text;
   699 
   700 
   701 (* refine *)
   702 
   703 fun gen_refine f text state =
   704   let
   705     val thy = Proof.theory_of state;
   706 
   707     fun eval (Basic mth) = f mth
   708       | eval (Source src) = f (method thy src)
   709       | eval (Then txts) = Seq.EVERY (map eval txts)
   710       | eval (Orelse txts) = Seq.FIRST (map eval txts)
   711       | eval (Try txt) = Seq.TRY (eval txt)
   712       | eval (Repeat1 txt) = Seq.REPEAT1 (eval txt);
   713   in eval text state end;
   714 
   715 val refine = gen_refine Proof.refine;
   716 val refine_end = gen_refine Proof.refine_end;
   717 
   718 
   719 (* structured proof steps *)
   720 
   721 val default_text = Source (Args.src (("default", []), Position.none));
   722 val this_text = Basic (K this);
   723 val done_text = Basic (K (SIMPLE_METHOD all_tac));
   724 
   725 fun close_text asm = Basic (fn ctxt => METHOD (K
   726   (FILTER Thm.no_prems ((if asm then ALLGOALS (assm_tac ctxt) else all_tac) THEN flexflex_tac))));
   727 
   728 fun finish_text asm None = close_text asm
   729   | finish_text asm (Some txt) = Then [txt, close_text asm];
   730 
   731 fun proof opt_text state =
   732   state
   733   |> Proof.assert_backward
   734   |> refine (if_none opt_text default_text)
   735   |> Seq.map (Proof.goal_facts (K []))
   736   |> Seq.map Proof.enter_forward;
   737 
   738 fun local_qed asm opt_text = Proof.local_qed (refine (finish_text asm opt_text));
   739 fun local_terminal_proof (text, opt_text) pr =
   740   Seq.THEN (proof (Some text), local_qed true opt_text pr);
   741 val local_default_proof = local_terminal_proof (default_text, None);
   742 val local_immediate_proof = local_terminal_proof (this_text, None);
   743 fun local_done_proof pr = Seq.THEN (proof (Some done_text), local_qed false None pr);
   744 
   745 
   746 fun global_qeds asm opt_text = Proof.global_qed (refine (finish_text asm opt_text));
   747 
   748 fun global_qed asm opt_text state =
   749   state
   750   |> global_qeds asm opt_text
   751   |> Proof.check_result "Failed to finish proof" state
   752   |> Seq.hd;
   753 
   754 fun global_term_proof asm (text, opt_text) state =
   755   state
   756   |> proof (Some text)
   757   |> Proof.check_result "Terminal proof method failed" state
   758   |> (Seq.flat o Seq.map (global_qeds asm opt_text))
   759   |> Proof.check_result "Failed to finish proof (after successful terminal method)" state
   760   |> Seq.hd;
   761 
   762 val global_terminal_proof = global_term_proof true;
   763 val global_default_proof = global_terminal_proof (default_text, None);
   764 val global_immediate_proof = global_terminal_proof (this_text, None);
   765 val global_done_proof = global_term_proof false (done_text, None);
   766 
   767 
   768 
   769 (** theory setup **)
   770 
   771 (* misc tactic emulations *)
   772 
   773 val subgoal_meth = goal_args_ctxt (Scan.repeat1 Args.name) subgoals_tac;
   774 val thin_meth = goal_args_ctxt Args.name thin_tac;
   775 val rename_meth = goal_args (Scan.repeat1 Args.name) Tactic.rename_params_tac;
   776 val rotate_meth = goal_args (Scan.optional Args.int 1) Tactic.rotate_tac;
   777 
   778 
   779 (* pure_methods *)
   780 
   781 val pure_methods =
   782  [("fail", no_args fail, "force failure"),
   783   ("succeed", no_args succeed, "succeed"),
   784   ("-", no_args insert_facts, "do nothing (insert current facts only)"),
   785   ("insert", thms_args insert, "insert theorems, ignoring facts (improper)"),
   786   ("unfold", thms_args unfold, "unfold definitions"),
   787   ("intro", thms_args intro, "repeatedly apply introduction rules"),
   788   ("elim", thms_args elim, "repeatedly apply elimination rules"),
   789   ("fold", thms_args fold, "fold definitions"),
   790   ("atomize", (atomize o #2) oo syntax (Args.mode "full"),
   791     "present local premises as object-level statements"),
   792   ("rules", rules_args rules_meth, "apply many rules, including proof search"),
   793   ("rule", thms_ctxt_args some_rule, "apply some intro/elim rule"),
   794   ("erule", nat_thms_args erule, "apply rule in elimination manner (improper)"),
   795   ("drule", nat_thms_args drule, "apply rule in destruct manner (improper)"),
   796   ("frule", nat_thms_args frule, "apply rule in forward manner (improper)"),
   797   ("this", no_args this, "apply current facts as rules"),
   798   ("assumption", ctxt_args assumption, "proof by assumption, preferring facts"),
   799   ("rule_tac", inst_args_var res_inst_meth, "apply rule (dynamic instantiation)"),
   800   ("erule_tac", inst_args_var eres_inst_meth, "apply rule in elimination manner (dynamic instantiation)"),
   801   ("drule_tac", inst_args_var dres_inst_meth, "apply rule in destruct manner (dynamic instantiation)"),
   802   ("frule_tac", inst_args_var forw_inst_meth, "apply rule in forward manner (dynamic instantiation)"),
   803   ("cut_tac", inst_args_var cut_inst_meth, "cut rule (dynamic instantiation)"),
   804   ("subgoal_tac", subgoal_meth, "insert subgoal (dynamic instantiation)"),
   805   ("thin_tac", thin_meth, "remove premise (dynamic instantiation)"),
   806   ("rename_tac", rename_meth, "rename parameters of goal (dynamic instantiation)"),
   807   ("rotate_tac", rotate_meth, "rotate assumptions of goal"),
   808   ("prolog", thms_args prolog, "simple prolog interpreter"),
   809   ("tactic", simple_args Args.name tactic, "ML tactic as proof method")];
   810 
   811 
   812 (* setup *)
   813 
   814 val setup =
   815  [MethodsData.init, add_methods pure_methods,
   816   (#1 o PureThy.add_thms [(("", Drule.equal_intr_rule), [ContextRules.intro_query_global None])])];
   817 
   818 
   819 end;
   820 
   821 
   822 structure BasicMethod: BASIC_METHOD = Method;
   823 open BasicMethod;