wenzelm@30160: (* Title: Tools/eqsubst.ML wenzelm@29269: Author: Lucas Dixon, University of Edinburgh paulson@15481: wenzelm@18598: A proof method to perform a substiution using an equation. wenzelm@18598: *) dixon@15538: wenzelm@18591: signature EQSUBST = paulson@15481: sig wenzelm@29269: (* a type abbreviation for match information *) dixon@19871: type match = dixon@19871: ((indexname * (sort * typ)) list (* type instantiations *) dixon@19871: * (indexname * (typ * term)) list) (* term instantiations *) dixon@19871: * (string * typ) list (* fake named type abs env *) dixon@19871: * (string * typ) list (* type abs env *) dixon@19871: * term (* outer term *) dixon@19871: dixon@19871: type searchinfo = dixon@19871: theory dixon@19871: * int (* maxidx *) dixon@19871: * 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@49339: val skip_first_asm_occs_search : wenzelm@49339: ('a -> 'b -> 'c Seq.seq Seq.seq) -> wenzelm@49339: 'a -> int -> 'b -> 'c skipseq wenzelm@49339: val skip_first_occs_search : wenzelm@49339: int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq wenzelm@49339: val skipto_skipseq : int -> 'a Seq.seq Seq.seq -> 'a skipseq dixon@19871: wenzelm@49339: (* tactics *) wenzelm@49339: val eqsubst_asm_tac : wenzelm@49339: Proof.context -> wenzelm@49339: int list -> thm list -> int -> tactic wenzelm@49339: val eqsubst_asm_tac' : wenzelm@49339: Proof.context -> wenzelm@49339: (searchinfo -> int -> term -> match skipseq) -> wenzelm@49339: int -> thm -> int -> tactic wenzelm@49339: val eqsubst_tac : wenzelm@49339: Proof.context -> wenzelm@49339: int list -> (* list of occurences to rewrite, use [0] for any *) wenzelm@49339: thm list -> int -> tactic wenzelm@49339: val eqsubst_tac' : wenzelm@49339: Proof.context -> (* proof context *) wenzelm@49339: (searchinfo -> term -> match Seq.seq) (* search function *) wenzelm@49339: -> thm (* equation theorem to rewrite with *) wenzelm@49339: -> int (* subgoal number in goal theorem *) wenzelm@49339: -> thm (* goal theorem *) wenzelm@49339: -> thm Seq.seq (* rewritten goal theorem *) dixon@19871: wenzelm@49339: (* search for substitutions *) wenzelm@49339: val valid_match_start : Zipper.T -> bool wenzelm@49339: val search_lr_all : Zipper.T -> Zipper.T Seq.seq wenzelm@49339: val search_lr_valid : (Zipper.T -> bool) -> Zipper.T -> Zipper.T Seq.seq wenzelm@49339: val searchf_lr_unify_all : wenzelm@49339: searchinfo -> term -> match Seq.seq Seq.seq wenzelm@49339: val searchf_lr_unify_valid : wenzelm@49339: searchinfo -> term -> match Seq.seq Seq.seq wenzelm@49339: val searchf_bt_unify_valid : wenzelm@49339: searchinfo -> term -> match Seq.seq Seq.seq dixon@19871: wenzelm@49339: val setup : theory -> theory dixon@19871: 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@32149: Simplifier.mksimps (simpset_of ctxt) #> map Drule.zero_var_indexes; wenzelm@18598: paulson@15481: dixon@15915: (* a type abriviation for match information *) wenzelm@16978: type match = wenzelm@16978: ((indexname * (sort * typ)) list (* type instantiations *) wenzelm@16978: * (indexname * (typ * term)) list) (* term instantiations *) wenzelm@16978: * (string * typ) list (* fake named type abs env *) wenzelm@16978: * (string * typ) list (* type abs env *) wenzelm@16978: * term (* outer term *) dixon@15550: wenzelm@16978: type searchinfo = wenzelm@18598: theory dixon@16004: * int (* maxidx *) dixon@19835: * 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. *) dixon@19835: datatype 'a skipseq = SkipMore of int dixon@19835: | SkipSeq of 'a Seq.seq Seq.seq; dixon@19835: (* given a seqseq, skip the first m non-empty seq's, note deficit *) wenzelm@49339: fun skipto_skipseq m s = wenzelm@49339: let wenzelm@49339: fun skip_occs n sq = wenzelm@49339: case Seq.pull sq of dixon@19835: NONE => SkipMore n wenzelm@49339: | SOME (h,t) => dixon@19835: (case Seq.pull h of NONE => skip_occs n t dixon@19835: | SOME _ => if n <= 1 then SkipSeq (Seq.cons h t) dixon@19835: else skip_occs (n - 1) t) dixon@19835: in (skip_occs m s) end; dixon@19835: wenzelm@49339: (* note: outerterm is the taget with the match replaced by a bound wenzelm@49339: variable : ie: "P lhs" beocmes "%x. P x" dixon@19835: insts is the types of instantiations of vars in lhs dixon@19835: and typinsts is the type instantiations of types in the lhs wenzelm@49339: Note: Final rule is the rule lifted into the ontext of the dixon@19835: taget thm. *) wenzelm@49339: fun mk_foo_match mkuptermfunc Ts t = wenzelm@49339: let dixon@19835: val ty = Term.type_of t dixon@19835: val bigtype = (rev (map snd Ts)) ---> ty dixon@19835: fun mk_foo 0 t = t dixon@19835: | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1))) dixon@19835: val num_of_bnds = (length Ts) dixon@19835: (* foo_term = "fooabs y0 ... yn" where y's are local bounds *) dixon@19835: val foo_term = mk_foo num_of_bnds (Bound num_of_bnds) dixon@19835: 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@49339: let val (FakeTs,Ts,newnames) = wenzelm@49339: List.foldr (fn ((n,ty),(FakeTs,Ts,usednames)) => wenzelm@43324: let val newname = singleton (Name.variant_list usednames) n wenzelm@49339: in ((mk_fake_bound_name newname,ty)::FakeTs, wenzelm@49339: (newname,ty)::Ts, dixon@19835: newname::usednames) end) dixon@19835: ([],[],[]) dixon@19835: Ts dixon@19835: 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 dixon@19835: abstraction. In this function we additionally construct the dixon@19835: abstraction environment, and an outer context term (with the focus dixon@19835: abstracted out) for use in rewriting with RWInst.rw *) wenzelm@49339: fun prep_zipper_match z = wenzelm@49339: let wenzelm@49339: val t = Zipper.trm z wenzelm@41164: val c = Zipper.ctxt z wenzelm@41164: val Ts = Zipper.C.nty_ctxt c dixon@19835: val (FakeTs', Ts', t') = fakefree_badbounds Ts t wenzelm@41164: val absterm = mk_foo_match (Zipper.C.apply c) Ts' t' dixon@19835: in dixon@19835: (t', (FakeTs', Ts', absterm)) dixon@19835: end; dixon@19835: wenzelm@49339: (* Unification with exception handled *) dixon@19835: (* given theory, max var index, pat, tgt; returns Seq of instantiations *) dixon@27033: fun clean_unify thry ix (a as (pat, tgt)) = dixon@19835: let dixon@19835: (* type info will be re-derived, maybe this can be cached dixon@19835: for efficiency? *) dixon@19835: val pat_ty = Term.type_of pat; dixon@19835: val tgt_ty = Term.type_of tgt; dixon@19835: (* is it OK to ignore the type instantiation info? dixon@19835: or should I be using it? *) dixon@19835: val typs_unify = wenzelm@29269: SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Vartab.empty, ix)) dixon@19835: handle Type.TUNIFY => NONE; dixon@19835: in dixon@19835: case typs_unify of dixon@19835: SOME (typinsttab, ix2) => dixon@19835: let dixon@19835: (* is it right to throw away the flexes? dixon@19835: 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}; dixon@27033: val useq = Unify.smash_unifiers thry [a] initenv wenzelm@40722: handle ListPair.UnequalLengths => Seq.empty wenzelm@32960: | Term.TERM _ => Seq.empty; dixon@19835: fun clean_unify' useq () = dixon@19835: (case (Seq.pull useq) of dixon@19835: NONE => NONE dixon@19835: | SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t))) wenzelm@40722: handle ListPair.UnequalLengths => NONE dixon@27033: | Term.TERM _ => NONE dixon@19835: in dixon@19835: (Seq.make (clean_unify' useq)) dixon@19835: end dixon@19835: | NONE => Seq.empty dixon@19835: end; dixon@19835: wenzelm@49339: (* Unification for zippers *) dixon@19835: (* Note: Ts is a modified version of the original names of the outer dixon@19835: bound variables. New names have been introduced to make sure they are dixon@19835: unique w.r.t all names in the term and each other. usednames' is dixon@19835: oldnames + new names. *) dixon@19835: (* ix = max var index *) wenzelm@49339: fun clean_unify_z sgn ix pat z = dixon@19835: let val (t, (FakeTs, Ts,absterm)) = prep_zipper_match z in wenzelm@49339: Seq.map (fn insts => (insts, FakeTs, Ts, absterm)) dixon@19835: (clean_unify sgn ix (t, pat)) end; dixon@19835: dixon@15550: dixon@19835: fun bot_left_leaf_of (l $ r) = bot_left_leaf_of l dixon@19835: | bot_left_leaf_of (Abs(s,ty,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@49339: (case bot_left_leaf_of (Zipper.trm z) of wenzelm@49339: Var _ => false dixon@19975: | _ => 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@49339: let wenzelm@49339: fun sf_valid_td_lr z = wenzelm@41164: let val here = if validf z then [Zipper.Here z] else [] in wenzelm@49339: case Zipper.trm z wenzelm@49339: of _ $ _ => [Zipper.LookIn (Zipper.move_down_left z)] wenzelm@49339: @ here wenzelm@41164: @ [Zipper.LookIn (Zipper.move_down_right z)] wenzelm@41164: | Abs _ => here @ [Zipper.LookIn (Zipper.move_down_abs z)] dixon@19835: | _ => here dixon@19835: end; wenzelm@41164: in Zipper.lzy_search sf_valid_td_lr end; dixon@15814: narboux@23064: (* search from bottom to top, left to right *) narboux@23064: narboux@23064: fun search_bt_valid validf = wenzelm@49339: let wenzelm@49339: fun sf_valid_td_lr z = wenzelm@41164: let val here = if validf z then [Zipper.Here z] else [] in wenzelm@49339: case Zipper.trm z wenzelm@49339: of _ $ _ => [Zipper.LookIn (Zipper.move_down_left z), wenzelm@41164: Zipper.LookIn (Zipper.move_down_right z)] @ here wenzelm@41164: | Abs _ => [Zipper.LookIn (Zipper.move_down_abs z)] @ here narboux@23064: | _ => here narboux@23064: end; wenzelm@41164: in Zipper.lzy_search sf_valid_td_lr end; narboux@23064: narboux@23064: fun searchf_unify_gen f (sgn, maxidx, z) lhs = wenzelm@49339: Seq.map (clean_unify_z sgn maxidx lhs) wenzelm@41164: (Zipper.limit_apply f z); narboux@23064: dixon@15814: (* search all unifications *) narboux@23064: val searchf_lr_unify_all = narboux@23064: searchf_unify_gen search_lr_all; paulson@15481: dixon@15814: (* search only for 'valid' unifiers (non abs subterms and non vars) *) wenzelm@49339: val searchf_lr_unify_valid = narboux@23064: searchf_unify_gen (search_lr_valid valid_match_start); dixon@15929: narboux@23064: val searchf_bt_unify_valid = narboux@23064: 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@49340: (RWInst.rw ctxt m rule conclthm) dixon@15855: |> IsaND.unfix_frees cfvs dixon@15915: |> RWInst.beta_eta_contract dixon@15538: |> (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@16978: let paulson@15481: val th = Thm.incr_indexes 1 gth; paulson@15481: val tgt_term = Thm.prop_of th; paulson@15481: wenzelm@22578: val sgn = Thm.theory_of_thm th; paulson@15481: val ctermify = Thm.cterm_of sgn; paulson@15481: val trivify = Thm.trivial o ctermify; paulson@15481: wenzelm@49340: val (fixedbody, fvs) = IsaND.fix_alls_term ctxt i tgt_term; paulson@15481: val cfvs = rev (map ctermify fvs); paulson@15481: dixon@15538: val conclterm = Logic.strip_imp_concl fixedbody; dixon@15538: val conclthm = trivify conclterm; dixon@27033: val maxidx = Thm.maxidx_of th; wenzelm@41164: val ft = ((Zipper.move_down_right (* ==> *) wenzelm@41164: o Zipper.move_down_left (* Trueprop *) wenzelm@41164: o Zipper.mktop dixon@16004: o Thm.prop_of) conclthm) paulson@15481: in dixon@16004: ((cfvs, conclthm), (sgn, maxidx, ft)) paulson@15481: end; paulson@15481: paulson@15481: (* substitute using an object or meta level equality *) wenzelm@18598: fun eqsubst_tac' ctxt searchf instepthm i th = wenzelm@16978: let wenzelm@49340: val (cvfsconclthm, searchinfo) = prep_concl_subst ctxt i th; wenzelm@18598: val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); dixon@15538: fun rewrite_with_thm r = dixon@15538: let val (lhs,_) = Logic.dest_equals (Thm.concl_of r); wenzelm@18598: in searchf searchinfo lhs wenzelm@49340: |> Seq.maps (apply_subst_in_concl ctxt i th cvfsconclthm r) end; wenzelm@18598: in stepthms |> Seq.maps rewrite_with_thm end; dixon@15538: dixon@15538: wenzelm@19047: (* distinct subgoals *) dixon@15959: fun distinct_subgoals th = wenzelm@19047: the_default th (SINGLE distinct_subgoals_tac th); dixon@15538: wenzelm@19047: (* General substitution of multiple occurances using one of dixon@15936: the given theorems*) dixon@19835: dixon@19835: wenzelm@16978: fun skip_first_occs_search occ srchf sinfo lhs = dixon@19835: case (skipto_skipseq occ (srchf sinfo lhs)) of dixon@19835: SkipMore _ => Seq.empty dixon@19835: | 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 = dixon@15936: let val nprems = Thm.nprems_of th in dixon@15936: if nprems < i then Seq.empty else wenzelm@16978: let val thmseq = (Seq.of_list thms) wenzelm@16978: fun apply_occ occ th = wenzelm@18598: thmseq |> Seq.maps wenzelm@49339: (fn r => eqsubst_tac' wenzelm@49339: ctxt dixon@19835: (skip_first_occs_search dixon@19871: occ searchf_lr_unify_valid) r dixon@15936: (i + ((Thm.nprems_of th) - nprems)) dixon@15936: th); wenzelm@16978: val sortedoccL = dixon@16004: Library.sort (Library.rev_order o Library.int_ord) occL; dixon@15936: in dixon@16004: Seq.map distinct_subgoals (Seq.EVERY (map apply_occ sortedoccL) th) dixon@15936: end wenzelm@49339: 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@18598: fun eqsubst_meth ctxt occL inthms = wenzelm@30510: SIMPLE_METHOD' (eqsubst_tac ctxt occL inthms); paulson@15481: dixon@16004: (* apply a substitution inside assumption j, keeps asm in the same place *) wenzelm@49340: fun apply_subst_in_asm ctxt i th rule ((cfvs, j, ngoalprems, pth),m) = wenzelm@16978: let dixon@16004: val th2 = Thm.rotate_rule (j - 1) i th; (* put premice first *) wenzelm@16978: val preelimrule = wenzelm@49340: (RWInst.rw ctxt m rule pth) wenzelm@21708: |> (Seq.hd o prune_params_tac) dixon@16004: |> Thm.permute_prems 0 ~1 (* put old asm first *) dixon@16004: |> IsaND.unfix_frees cfvs (* unfix any global params *) dixon@16004: |> RWInst.beta_eta_contract; (* normal form *) wenzelm@16978: (* val elimrule = dixon@16004: preelimrule dixon@16004: |> Tactic.make_elim (* make into elim rule *) dixon@16004: |> Thm.lift_rule (th2, i); (* lift into context *) dixon@16007: *) dixon@16004: in dixon@16004: (* ~j because new asm starts at back, thus we subtract 1 *) dixon@16007: Seq.map (Thm.rotate_rule (~j) ((Thm.nprems_of rule) + i)) dixon@16007: (Tactic.dtac preelimrule i th2) dixon@16007: wenzelm@16978: (* (Thm.bicompose dixon@16004: false (* use unification *) dixon@16004: (true, (* elim resolution *) dixon@16007: elimrule, (2 + (Thm.nprems_of rule)) - ngoalprems) dixon@16007: i th2) *) dixon@16004: 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@16978: let paulson@15481: val th = Thm.incr_indexes 1 gth; paulson@15481: val tgt_term = Thm.prop_of th; paulson@15481: wenzelm@22578: val sgn = Thm.theory_of_thm th; paulson@15481: val ctermify = Thm.cterm_of sgn; paulson@15481: val trivify = Thm.trivial o ctermify; paulson@15481: wenzelm@49340: val (fixedbody, fvs) = IsaND.fix_alls_term ctxt i tgt_term; paulson@15481: val cfvs = rev (map ctermify fvs); paulson@15481: haftmann@18011: val asmt = nth (Logic.strip_imp_prems fixedbody) (j - 1); dixon@15538: val asm_nprems = length (Logic.strip_imp_prems asmt); dixon@15538: dixon@15538: val pth = trivify asmt; dixon@27033: val maxidx = Thm.maxidx_of th; dixon@15538: wenzelm@41164: val ft = ((Zipper.move_down_right (* trueprop *) wenzelm@41164: o Zipper.mktop dixon@16004: o Thm.prop_of) pth) dixon@16004: in ((cfvs, j, asm_nprems, pth), (sgn, maxidx, ft)) end; paulson@15481: dixon@15538: (* prepare subst in every possible assumption *) wenzelm@49340: fun prep_subst_in_asms ctxt i gth = wenzelm@49340: map (prep_subst_in_asm ctxt i gth) dixon@19835: ((fn l => Library.upto (1, length l)) dixon@16004: (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@16978: let wenzelm@49340: val asmpreps = prep_subst_in_asms ctxt i th; wenzelm@18598: val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); dixon@16004: fun rewrite_with_thm r = dixon@16004: let val (lhs,_) = Logic.dest_equals (Thm.concl_of r) dixon@16004: fun occ_search occ [] = Seq.empty dixon@16004: | occ_search occ ((asminfo, searchinfo)::moreasms) = wenzelm@16978: (case searchf searchinfo occ lhs of dixon@19835: SkipMore i => occ_search i moreasms dixon@19835: | SkipSeq ss => wenzelm@19861: Seq.append (Seq.map (Library.pair asminfo) (Seq.flat ss)) wenzelm@19861: (occ_search 1 moreasms)) dixon@16004: (* find later substs also *) wenzelm@16978: in wenzelm@49340: occ_search skipocc asmpreps |> Seq.maps (apply_subst_in_asm ctxt i th r) dixon@16004: end; wenzelm@18598: in stepthms |> Seq.maps rewrite_with_thm end; dixon@15538: dixon@16004: wenzelm@16978: fun skip_first_asm_occs_search searchf sinfo occ lhs = dixon@19835: skipto_skipseq occ (searchf sinfo lhs); dixon@16004: wenzelm@18598: fun eqsubst_asm_tac ctxt occL thms i th = wenzelm@16978: let val nprems = Thm.nprems_of th dixon@15538: in dixon@16004: if nprems < i then Seq.empty else wenzelm@16978: let val thmseq = (Seq.of_list thms) wenzelm@16978: fun apply_occ occK th = wenzelm@18598: thmseq |> Seq.maps wenzelm@16978: (fn r => wenzelm@18598: eqsubst_asm_tac' ctxt (skip_first_asm_occs_search dixon@19871: searchf_lr_unify_valid) occK r dixon@16004: (i + ((Thm.nprems_of th) - nprems)) dixon@16004: th); wenzelm@16978: val sortedoccs = dixon@16004: Library.sort (Library.rev_order o Library.int_ord) occL dixon@16004: in dixon@16004: Seq.map distinct_subgoals dixon@16004: (Seq.EVERY (map apply_occ sortedoccs) th) dixon@16004: end wenzelm@49339: 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@30510: 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;