wenzelm@30160: (* Title: Tools/eqsubst.ML wenzelm@29269: Author: Lucas Dixon, University of Edinburgh paulson@15481: wenzelm@52235: Perform a substitution using an equation. wenzelm@18598: *) dixon@15538: wenzelm@18591: signature EQSUBST = paulson@15481: sig dixon@19871: type match = wenzelm@52234: ((indexname * (sort * typ)) list (* type instantiations *) wenzelm@52234: * (indexname * (typ * term)) list) (* term instantiations *) wenzelm@52234: * (string * typ) list (* fake named type abs env *) wenzelm@52234: * (string * typ) list (* type abs env *) wenzelm@52234: * term (* outer term *) dixon@19871: dixon@19871: type searchinfo = wenzelm@52234: theory wenzelm@52234: * int (* maxidx *) wenzelm@52234: * Zipper.T (* focusterm to search under *) dixon@19871: wenzelm@49339: datatype 'a skipseq = SkipMore of int | SkipSeq of 'a Seq.seq Seq.seq dixon@19871: wenzelm@52234: val skip_first_asm_occs_search: ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> int -> 'b -> 'c skipseq wenzelm@52234: val skip_first_occs_search: int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq wenzelm@52234: val skipto_skipseq: int -> 'a Seq.seq Seq.seq -> 'a skipseq dixon@19871: wenzelm@49339: (* tactics *) wenzelm@52234: val eqsubst_asm_tac: Proof.context -> int list -> thm list -> int -> tactic wenzelm@52234: val eqsubst_asm_tac': Proof.context -> wenzelm@52234: (searchinfo -> int -> term -> match skipseq) -> int -> thm -> int -> tactic wenzelm@52234: val eqsubst_tac: Proof.context -> wenzelm@52234: int list -> (* list of occurences to rewrite, use [0] for any *) wenzelm@52234: thm list -> int -> tactic wenzelm@52234: val eqsubst_tac': Proof.context -> wenzelm@52234: (searchinfo -> term -> match Seq.seq) (* search function *) wenzelm@52234: -> thm (* equation theorem to rewrite with *) wenzelm@52234: -> int (* subgoal number in goal theorem *) wenzelm@52234: -> thm (* goal theorem *) wenzelm@52234: -> thm Seq.seq (* rewritten goal theorem *) dixon@19871: wenzelm@49339: (* search for substitutions *) wenzelm@52234: val valid_match_start: Zipper.T -> bool wenzelm@52234: val search_lr_all: Zipper.T -> Zipper.T Seq.seq wenzelm@52234: val search_lr_valid: (Zipper.T -> bool) -> Zipper.T -> Zipper.T Seq.seq wenzelm@52234: val searchf_lr_unify_all: searchinfo -> term -> match Seq.seq Seq.seq wenzelm@52234: val searchf_lr_unify_valid: searchinfo -> term -> match Seq.seq Seq.seq wenzelm@52234: val searchf_bt_unify_valid: searchinfo -> term -> match Seq.seq Seq.seq dixon@19871: wenzelm@49339: val setup : theory -> theory paulson@15481: end; paulson@15481: wenzelm@41164: structure EqSubst: EQSUBST = wenzelm@41164: struct dixon@19835: dixon@19835: (* changes object "=" to meta "==" which prepares a given rewrite rule *) wenzelm@18598: fun prep_meta_eq ctxt = wenzelm@51717: Simplifier.mksimps ctxt #> map Drule.zero_var_indexes; wenzelm@18598: paulson@15481: wenzelm@52234: type match = wenzelm@52235: ((indexname * (sort * typ)) list (* type instantiations *) wenzelm@52235: * (indexname * (typ * term)) list) (* term instantiations *) wenzelm@52235: * (string * typ) list (* fake named type abs env *) wenzelm@52235: * (string * typ) list (* type abs env *) wenzelm@52235: * term; (* outer term *) dixon@15550: wenzelm@52234: type searchinfo = wenzelm@52235: theory wenzelm@52235: * int (* maxidx *) wenzelm@52235: * Zipper.T; (* focusterm to search under *) dixon@19835: dixon@19835: dixon@19835: (* skipping non-empty sub-sequences but when we reach the end dixon@19835: of the seq, remembering how much we have left to skip. *) wenzelm@52234: datatype 'a skipseq = wenzelm@52234: SkipMore of int | wenzelm@52234: SkipSeq of 'a Seq.seq Seq.seq; wenzelm@52234: dixon@19835: (* given a seqseq, skip the first m non-empty seq's, note deficit *) wenzelm@49339: fun skipto_skipseq m s = wenzelm@52234: let wenzelm@52234: fun skip_occs n sq = wenzelm@52234: (case Seq.pull sq of wenzelm@52234: NONE => SkipMore n wenzelm@52234: | SOME (h,t) => wenzelm@52234: (case Seq.pull h of wenzelm@52234: NONE => skip_occs n t wenzelm@52234: | SOME _ => if n <= 1 then SkipSeq (Seq.cons h t) else skip_occs (n - 1) t)) wenzelm@52234: in skip_occs m s end; dixon@19835: wenzelm@49339: (* note: outerterm is the taget with the match replaced by a bound wenzelm@52234: variable : ie: "P lhs" beocmes "%x. P x" wenzelm@52234: insts is the types of instantiations of vars in lhs wenzelm@52234: and typinsts is the type instantiations of types in the lhs wenzelm@52234: Note: Final rule is the rule lifted into the ontext of the wenzelm@52234: taget thm. *) wenzelm@49339: fun mk_foo_match mkuptermfunc Ts t = wenzelm@52234: let wenzelm@52234: val ty = Term.type_of t wenzelm@52235: val bigtype = rev (map snd Ts) ---> ty wenzelm@52234: fun mk_foo 0 t = t wenzelm@52234: | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1))) wenzelm@52235: val num_of_bnds = length Ts wenzelm@52234: (* foo_term = "fooabs y0 ... yn" where y's are local bounds *) wenzelm@52234: val foo_term = mk_foo num_of_bnds (Bound num_of_bnds) wenzelm@52234: in Abs ("fooabs", bigtype, mkuptermfunc foo_term) end; dixon@19835: dixon@19835: (* T is outer bound vars, n is number of locally bound vars *) dixon@19835: (* THINK: is order of Ts correct...? or reversed? *) wenzelm@49339: fun mk_fake_bound_name n = ":b_" ^ n; wenzelm@49339: fun fakefree_badbounds Ts t = wenzelm@52234: let val (FakeTs, Ts, newnames) = wenzelm@52234: List.foldr (fn ((n, ty), (FakeTs, Ts, usednames)) => wenzelm@52234: let wenzelm@52234: val newname = singleton (Name.variant_list usednames) n wenzelm@52234: in wenzelm@52234: ((mk_fake_bound_name newname, ty) :: FakeTs, wenzelm@52234: (newname, ty) :: Ts, wenzelm@52234: newname :: usednames) wenzelm@52234: end) ([], [], []) Ts wenzelm@52234: in (FakeTs, Ts, Term.subst_bounds (map Free FakeTs, t)) end; dixon@19835: dixon@19835: (* before matching we need to fake the bound vars that are missing an wenzelm@52235: abstraction. In this function we additionally construct the wenzelm@52235: abstraction environment, and an outer context term (with the focus wenzelm@52235: abstracted out) for use in rewriting with RWInst.rw *) wenzelm@49339: fun prep_zipper_match z = wenzelm@52234: let wenzelm@52234: val t = Zipper.trm z wenzelm@52234: val c = Zipper.ctxt z wenzelm@52234: val Ts = Zipper.C.nty_ctxt c wenzelm@52234: val (FakeTs', Ts', t') = fakefree_badbounds Ts t wenzelm@52234: val absterm = mk_foo_match (Zipper.C.apply c) Ts' t' wenzelm@52234: in wenzelm@52234: (t', (FakeTs', Ts', absterm)) wenzelm@52234: end; dixon@19835: wenzelm@49339: (* Unification with exception handled *) dixon@19835: (* given theory, max var index, pat, tgt; returns Seq of instantiations *) wenzelm@52235: fun clean_unify thy ix (a as (pat, tgt)) = wenzelm@52234: let wenzelm@52234: (* type info will be re-derived, maybe this can be cached wenzelm@52234: for efficiency? *) wenzelm@52234: val pat_ty = Term.type_of pat; wenzelm@52234: val tgt_ty = Term.type_of tgt; wenzelm@52235: (* FIXME is it OK to ignore the type instantiation info? wenzelm@52234: or should I be using it? *) wenzelm@52234: val typs_unify = wenzelm@52235: SOME (Sign.typ_unify thy (pat_ty, tgt_ty) (Vartab.empty, ix)) wenzelm@52234: handle Type.TUNIFY => NONE; wenzelm@52234: in wenzelm@52234: (case typs_unify of wenzelm@52234: SOME (typinsttab, ix2) => dixon@19835: let wenzelm@52234: (* FIXME is it right to throw away the flexes? wenzelm@52234: or should I be using them somehow? *) dixon@19835: fun mk_insts env = dixon@19835: (Vartab.dest (Envir.type_env env), wenzelm@32032: Vartab.dest (Envir.term_env env)); wenzelm@32032: val initenv = wenzelm@32032: Envir.Envir {maxidx = ix2, tenv = Vartab.empty, tyenv = typinsttab}; wenzelm@52235: val useq = Unify.smash_unifiers thy [a] initenv wenzelm@52234: handle ListPair.UnequalLengths => Seq.empty wenzelm@52234: | Term.TERM _ => Seq.empty; dixon@19835: fun clean_unify' useq () = wenzelm@52234: (case (Seq.pull useq) of wenzelm@52234: NONE => NONE wenzelm@52234: | SOME (h, t) => SOME (mk_insts h, Seq.make (clean_unify' t))) wenzelm@52234: handle ListPair.UnequalLengths => NONE wenzelm@52234: | Term.TERM _ => NONE; dixon@19835: in dixon@19835: (Seq.make (clean_unify' useq)) dixon@19835: end wenzelm@52234: | NONE => Seq.empty) wenzelm@52234: end; dixon@19835: wenzelm@49339: (* Unification for zippers *) dixon@19835: (* Note: Ts is a modified version of the original names of the outer wenzelm@52235: bound variables. New names have been introduced to make sure they are wenzelm@52235: unique w.r.t all names in the term and each other. usednames' is wenzelm@52235: oldnames + new names. *) wenzelm@52235: fun clean_unify_z thy maxidx pat z = wenzelm@52235: let val (t, (FakeTs, Ts, absterm)) = prep_zipper_match z in wenzelm@49339: Seq.map (fn insts => (insts, FakeTs, Ts, absterm)) wenzelm@52235: (clean_unify thy maxidx (t, pat)) wenzelm@52234: end; dixon@19835: dixon@15550: wenzelm@52234: fun bot_left_leaf_of (l $ _) = bot_left_leaf_of l wenzelm@52234: | bot_left_leaf_of (Abs (_, _, t)) = bot_left_leaf_of t dixon@19835: | bot_left_leaf_of x = x; dixon@15538: dixon@19975: (* Avoid considering replacing terms which have a var at the head as dixon@19975: they always succeed trivially, and uninterestingly. *) dixon@19835: fun valid_match_start z = wenzelm@52234: (case bot_left_leaf_of (Zipper.trm z) of wenzelm@52234: Var _ => false wenzelm@52234: | _ => true); dixon@19975: dixon@15814: (* search from top, left to right, then down *) dixon@19871: val search_lr_all = ZipperSearch.all_bl_ur; paulson@15481: dixon@15814: (* search from top, left to right, then down *) dixon@19871: fun search_lr_valid validf = wenzelm@52234: let wenzelm@52234: fun sf_valid_td_lr z = wenzelm@52234: let val here = if validf z then [Zipper.Here z] else [] in wenzelm@52234: (case Zipper.trm z of wenzelm@52234: _ $ _ => wenzelm@52234: [Zipper.LookIn (Zipper.move_down_left z)] @ here @ wenzelm@52234: [Zipper.LookIn (Zipper.move_down_right z)] wenzelm@52234: | Abs _ => here @ [Zipper.LookIn (Zipper.move_down_abs z)] wenzelm@52234: | _ => here) wenzelm@52234: end; wenzelm@52234: in Zipper.lzy_search sf_valid_td_lr end; dixon@15814: narboux@23064: (* search from bottom to top, left to right *) narboux@23064: fun search_bt_valid validf = wenzelm@52234: let wenzelm@52234: fun sf_valid_td_lr z = wenzelm@52234: let val here = if validf z then [Zipper.Here z] else [] in wenzelm@52234: (case Zipper.trm z of wenzelm@52234: _ $ _ => wenzelm@52234: [Zipper.LookIn (Zipper.move_down_left z), wenzelm@52234: Zipper.LookIn (Zipper.move_down_right z)] @ here wenzelm@52234: | Abs _ => [Zipper.LookIn (Zipper.move_down_abs z)] @ here wenzelm@52234: | _ => here) wenzelm@52234: end; wenzelm@52234: in Zipper.lzy_search sf_valid_td_lr end; narboux@23064: wenzelm@52235: fun searchf_unify_gen f (thy, maxidx, z) lhs = wenzelm@52235: Seq.map (clean_unify_z thy maxidx lhs) (Zipper.limit_apply f z); narboux@23064: dixon@15814: (* search all unifications *) wenzelm@52234: val searchf_lr_unify_all = searchf_unify_gen search_lr_all; paulson@15481: dixon@15814: (* search only for 'valid' unifiers (non abs subterms and non vars) *) wenzelm@52234: val searchf_lr_unify_valid = searchf_unify_gen (search_lr_valid valid_match_start); dixon@15929: wenzelm@52234: val searchf_bt_unify_valid = searchf_unify_gen (search_bt_valid valid_match_start); dixon@15814: dixon@15538: (* apply a substitution in the conclusion of the theorem th *) dixon@15538: (* cfvs are certified free var placeholders for goal params *) dixon@15538: (* conclthm is a theorem of for just the conclusion *) dixon@15538: (* m is instantiation/match information *) dixon@15538: (* rule is the equation for substitution *) wenzelm@49340: fun apply_subst_in_concl ctxt i th (cfvs, conclthm) rule m = wenzelm@52234: RWInst.rw ctxt m rule conclthm wenzelm@52234: |> IsaND.unfix_frees cfvs wenzelm@52234: |> RWInst.beta_eta_contract wenzelm@52234: |> (fn r => Tactic.rtac r i th); paulson@15481: paulson@15481: (* substitute within the conclusion of goal i of gth, using a meta dixon@15538: equation rule. Note that we assume rule has var indicies zero'd *) wenzelm@49340: fun prep_concl_subst ctxt i gth = wenzelm@52234: let wenzelm@52234: val th = Thm.incr_indexes 1 gth; wenzelm@52234: val tgt_term = Thm.prop_of th; paulson@15481: wenzelm@52235: val thy = Thm.theory_of_thm th; wenzelm@52235: val cert = Thm.cterm_of thy; paulson@15481: wenzelm@52234: val (fixedbody, fvs) = IsaND.fix_alls_term ctxt i tgt_term; wenzelm@52235: val cfvs = rev (map cert fvs); paulson@15481: wenzelm@52234: val conclterm = Logic.strip_imp_concl fixedbody; wenzelm@52235: val conclthm = Thm.trivial (cert conclterm); wenzelm@52234: val maxidx = Thm.maxidx_of th; wenzelm@52234: val ft = wenzelm@52234: (Zipper.move_down_right (* ==> *) wenzelm@52234: o Zipper.move_down_left (* Trueprop *) wenzelm@52234: o Zipper.mktop wenzelm@52234: o Thm.prop_of) conclthm wenzelm@52234: in wenzelm@52235: ((cfvs, conclthm), (thy, maxidx, ft)) wenzelm@52234: end; paulson@15481: paulson@15481: (* substitute using an object or meta level equality *) wenzelm@18598: fun eqsubst_tac' ctxt searchf instepthm i th = wenzelm@52234: let wenzelm@52234: val (cvfsconclthm, searchinfo) = prep_concl_subst ctxt i th; wenzelm@52234: val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); wenzelm@52234: fun rewrite_with_thm r = wenzelm@52234: let val (lhs,_) = Logic.dest_equals (Thm.concl_of r) in wenzelm@52234: searchf searchinfo lhs wenzelm@52234: |> Seq.maps (apply_subst_in_concl ctxt i th cvfsconclthm r) wenzelm@52234: end; wenzelm@52234: in stepthms |> Seq.maps rewrite_with_thm end; dixon@15538: dixon@15538: wenzelm@19047: (* distinct subgoals *) wenzelm@52234: fun distinct_subgoals th = the_default th (SINGLE distinct_subgoals_tac th); wenzelm@52234: dixon@15538: wenzelm@19047: (* General substitution of multiple occurances using one of wenzelm@52235: the given theorems *) dixon@19835: wenzelm@16978: fun skip_first_occs_search occ srchf sinfo lhs = wenzelm@52234: (case (skipto_skipseq occ (srchf sinfo lhs)) of wenzelm@52234: SkipMore _ => Seq.empty wenzelm@52234: | SkipSeq ss => Seq.flat ss); dixon@16004: dixon@22727: (* The occL is a list of integers indicating which occurence dixon@22727: w.r.t. the search order, to rewrite. Backtracking will also find later dixon@22727: occurences, but all earlier ones are skipped. Thus you can use [0] to dixon@22727: just find all rewrites. *) dixon@22727: wenzelm@18598: fun eqsubst_tac ctxt occL thms i th = wenzelm@52234: let val nprems = Thm.nprems_of th in wenzelm@52234: if nprems < i then Seq.empty else wenzelm@52234: let wenzelm@52234: val thmseq = (Seq.of_list thms); wenzelm@52234: fun apply_occ occ th = wenzelm@52234: thmseq |> Seq.maps (fn r => wenzelm@52234: eqsubst_tac' ctxt wenzelm@52234: (skip_first_occs_search occ searchf_lr_unify_valid) r wenzelm@52235: (i + (Thm.nprems_of th - nprems)) th); wenzelm@52234: val sortedoccL = Library.sort (rev_order o int_ord) occL; wenzelm@52234: in wenzelm@52234: Seq.map distinct_subgoals (Seq.EVERY (map apply_occ sortedoccL) th) wenzelm@52234: end wenzelm@52234: end; dixon@15959: paulson@15481: paulson@15481: (* inthms are the given arguments in Isar, and treated as eqstep with paulson@15481: the first one, then the second etc *) wenzelm@52234: fun eqsubst_meth ctxt occL inthms = SIMPLE_METHOD' (eqsubst_tac ctxt occL inthms); paulson@15481: dixon@16004: (* apply a substitution inside assumption j, keeps asm in the same place *) wenzelm@52234: fun apply_subst_in_asm ctxt i th rule ((cfvs, j, _, pth),m) = wenzelm@52234: let wenzelm@52234: val th2 = Thm.rotate_rule (j - 1) i th; (* put premice first *) wenzelm@52234: val preelimrule = wenzelm@52234: RWInst.rw ctxt m rule pth wenzelm@52234: |> (Seq.hd o prune_params_tac) wenzelm@52234: |> Thm.permute_prems 0 ~1 (* put old asm first *) wenzelm@52234: |> IsaND.unfix_frees cfvs (* unfix any global params *) wenzelm@52234: |> RWInst.beta_eta_contract; (* normal form *) wenzelm@52234: in wenzelm@52234: (* ~j because new asm starts at back, thus we subtract 1 *) wenzelm@52235: Seq.map (Thm.rotate_rule (~ j) (Thm.nprems_of rule + i)) dixon@16007: (Tactic.dtac preelimrule i th2) wenzelm@52234: end; paulson@15481: paulson@15481: dixon@15538: (* prepare to substitute within the j'th premise of subgoal i of gth, dixon@15538: using a meta-level equation. Note that we assume rule has var indicies dixon@15538: zero'd. Note that we also assume that premt is the j'th premice of dixon@15538: subgoal i of gth. Note the repetition of work done for each dixon@15538: assumption, i.e. this can be made more efficient for search over dixon@15538: multiple assumptions. *) wenzelm@49340: fun prep_subst_in_asm ctxt i gth j = wenzelm@52234: let wenzelm@52234: val th = Thm.incr_indexes 1 gth; wenzelm@52234: val tgt_term = Thm.prop_of th; paulson@15481: wenzelm@52235: val thy = Thm.theory_of_thm th; wenzelm@52235: val cert = Thm.cterm_of thy; paulson@15481: wenzelm@52234: val (fixedbody, fvs) = IsaND.fix_alls_term ctxt i tgt_term; wenzelm@52235: val cfvs = rev (map cert fvs); paulson@15481: wenzelm@52234: val asmt = nth (Logic.strip_imp_prems fixedbody) (j - 1); wenzelm@52234: val asm_nprems = length (Logic.strip_imp_prems asmt); wenzelm@52234: wenzelm@52235: val pth = Thm.trivial (cert asmt); wenzelm@52234: val maxidx = Thm.maxidx_of th; dixon@15538: wenzelm@52234: val ft = wenzelm@52234: (Zipper.move_down_right (* trueprop *) wenzelm@52234: o Zipper.mktop wenzelm@52234: o Thm.prop_of) pth wenzelm@52235: in ((cfvs, j, asm_nprems, pth), (thy, maxidx, ft)) end; paulson@15481: dixon@15538: (* prepare subst in every possible assumption *) wenzelm@49340: fun prep_subst_in_asms ctxt i gth = wenzelm@52234: map (prep_subst_in_asm ctxt i gth) wenzelm@52234: ((fn l => Library.upto (1, length l)) wenzelm@52234: (Logic.prems_of_goal (Thm.prop_of gth) i)); dixon@15538: dixon@15538: dixon@15538: (* substitute in an assumption using an object or meta level equality *) wenzelm@18598: fun eqsubst_asm_tac' ctxt searchf skipocc instepthm i th = wenzelm@52234: let wenzelm@52234: val asmpreps = prep_subst_in_asms ctxt i th; wenzelm@52234: val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); wenzelm@52234: fun rewrite_with_thm r = wenzelm@52234: let wenzelm@52234: val (lhs,_) = Logic.dest_equals (Thm.concl_of r); wenzelm@52234: fun occ_search occ [] = Seq.empty wenzelm@52234: | occ_search occ ((asminfo, searchinfo)::moreasms) = wenzelm@52234: (case searchf searchinfo occ lhs of wenzelm@52234: SkipMore i => occ_search i moreasms wenzelm@52234: | SkipSeq ss => wenzelm@52234: Seq.append (Seq.map (Library.pair asminfo) (Seq.flat ss)) wenzelm@52234: (occ_search 1 moreasms)) (* find later substs also *) wenzelm@52234: in wenzelm@52234: occ_search skipocc asmpreps |> Seq.maps (apply_subst_in_asm ctxt i th r) wenzelm@52234: end; wenzelm@52234: in stepthms |> Seq.maps rewrite_with_thm end; dixon@15538: dixon@16004: wenzelm@16978: fun skip_first_asm_occs_search searchf sinfo occ lhs = wenzelm@52234: skipto_skipseq occ (searchf sinfo lhs); dixon@16004: wenzelm@18598: fun eqsubst_asm_tac ctxt occL thms i th = wenzelm@52234: let val nprems = Thm.nprems_of th in wenzelm@52234: if nprems < i then Seq.empty wenzelm@52234: else wenzelm@52234: let wenzelm@52234: val thmseq = Seq.of_list thms; wenzelm@16978: fun apply_occ occK th = wenzelm@52234: thmseq |> Seq.maps (fn r => wenzelm@52234: eqsubst_asm_tac' ctxt wenzelm@52234: (skip_first_asm_occs_search searchf_lr_unify_valid) occK r wenzelm@52235: (i + (Thm.nprems_of th - nprems)) th); wenzelm@52234: val sortedoccs = Library.sort (rev_order o int_ord) occL; dixon@16004: in wenzelm@52234: Seq.map distinct_subgoals (Seq.EVERY (map apply_occ sortedoccs) th) dixon@16004: end wenzelm@52234: end; paulson@15481: paulson@15481: (* inthms are the given arguments in Isar, and treated as eqstep with paulson@15481: the first one, then the second etc *) wenzelm@18598: fun eqsubst_asm_meth ctxt occL inthms = wenzelm@52234: SIMPLE_METHOD' (eqsubst_asm_tac ctxt occL inthms); paulson@15481: wenzelm@18598: (* combination method that takes a flag (true indicates that subst wenzelm@31301: should be done to an assumption, false = apply to the conclusion of wenzelm@31301: the goal) as well as the theorems to use *) wenzelm@16978: val setup = wenzelm@31301: Method.setup @{binding subst} wenzelm@44095: (Args.mode "asm" -- Scan.lift (Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0]) -- wenzelm@44095: Attrib.thms >> wenzelm@44095: (fn ((asm, occL), inthms) => fn ctxt => wenzelm@44095: (if asm then eqsubst_asm_meth else eqsubst_meth) ctxt occL inthms)) wenzelm@31301: "single-step substitution"; paulson@15481: wenzelm@16978: end;