paulson@15481: (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) dixon@15538: (* Title: Provers/eqsubst.ML paulson@15481: Author: Lucas Dixon, University of Edinburgh paulson@15481: lucas.dixon@ed.ac.uk dixon@15538: Modified: 18 Feb 2005 - Lucas - paulson@15481: Created: 29 Jan 2005 paulson@15481: *) paulson@15481: (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) paulson@15481: (* DESCRIPTION: paulson@15481: paulson@15481: A Tactic to perform a substiution using an equation. paulson@15481: paulson@15481: *) paulson@15481: (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) paulson@15481: dixon@16004: dixon@16004: dixon@15538: (* Logic specific data stub *) paulson@15481: signature EQRULE_DATA = paulson@15481: sig dixon@15538: paulson@15481: (* to make a meta equality theorem in the current logic *) paulson@15481: val prep_meta_eq : thm -> thm list dixon@15538: paulson@15481: end; paulson@15481: dixon@15538: paulson@15481: (* the signature of an instance of the SQSUBST tactic *) paulson@15481: signature EQSUBST_TAC = paulson@15481: sig dixon@15538: dixon@15959: exception eqsubst_occL_exp of dixon@15959: string * (int list) * (Thm.thm list) * int * Thm.thm; dixon@15959: dixon@16004: dixon@15550: type match = dixon@15915: ((Term.indexname * (Term.sort * Term.typ)) list (* type instantiations *) dixon@15915: * (Term.indexname * (Term.typ * Term.term)) list) (* term instantiations *) dixon@15814: * (string * Term.typ) list (* fake named type abs env *) dixon@15814: * (string * Term.typ) list (* type abs env *) dixon@15814: * Term.term (* outer term *) dixon@15550: dixon@16004: type searchinfo = dixon@16004: Sign.sg (* sign for matching *) dixon@16004: * int (* maxidx *) dixon@16004: * BasicIsaFTerm.FcTerm (* focusterm to search under *) dixon@16004: dixon@15538: val prep_subst_in_asm : dixon@16004: int (* subgoal to subst in *) dixon@15538: -> Thm.thm (* target theorem with subgoals *) dixon@15538: -> int (* premise to subst in *) dixon@15538: -> (Thm.cterm list (* certified free var placeholders for vars *) dixon@15538: * int (* premice no. to subst *) dixon@15538: * int (* number of assumptions of premice *) dixon@15538: * Thm.thm) (* premice as a new theorem for forward reasoning *) dixon@16004: * searchinfo (* search info: prem id etc *) dixon@15538: dixon@15538: val prep_subst_in_asms : dixon@16004: int (* subgoal to subst in *) dixon@15538: -> Thm.thm (* target theorem with subgoals *) dixon@15538: -> ((Thm.cterm list (* certified free var placeholders for vars *) dixon@15538: * int (* premice no. to subst *) dixon@15538: * int (* number of assumptions of premice *) dixon@15538: * Thm.thm) (* premice as a new theorem for forward reasoning *) dixon@16004: * searchinfo) list dixon@15538: dixon@15538: val apply_subst_in_asm : dixon@15538: int (* subgoal *) dixon@15538: -> Thm.thm (* overall theorem *) dixon@16004: -> Thm.thm (* rule *) dixon@15538: -> (Thm.cterm list (* certified free var placeholders for vars *) dixon@15538: * int (* assump no being subst *) dixon@15538: * int (* num of premises of asm *) dixon@15538: * Thm.thm) (* premthm *) dixon@16004: * match dixon@15538: -> Thm.thm Seq.seq dixon@15538: dixon@15538: val prep_concl_subst : dixon@16004: int (* subgoal *) dixon@15538: -> Thm.thm (* overall goal theorem *) dixon@16004: -> (Thm.cterm list * Thm.thm) * searchinfo (* (cvfs, conclthm), matchf *) dixon@15538: dixon@15538: val apply_subst_in_concl : dixon@15538: int (* subgoal *) dixon@15538: -> Thm.thm (* thm with all goals *) dixon@15538: -> Thm.cterm list (* certified free var placeholders for vars *) dixon@15538: * Thm.thm (* trivial thm of goal concl *) dixon@15538: (* possible matches/unifiers *) dixon@15538: -> Thm.thm (* rule *) dixon@15550: -> match dixon@15538: -> Thm.thm Seq.seq (* substituted goal *) dixon@15538: dixon@16004: (* basic notion of search *) dixon@15814: val searchf_tlr_unify_all : dixon@16004: (searchinfo dixon@16004: -> Term.term (* lhs *) dixon@16004: -> match Seq.seq Seq.seq) dixon@15814: val searchf_tlr_unify_valid : dixon@16004: (searchinfo dixon@16004: -> Term.term (* lhs *) dixon@16004: -> match Seq.seq Seq.seq) dixon@15814: dixon@16004: (* specialise search constructor for conclusion skipping occurnaces. *) dixon@16004: val skip_first_occs_search : dixon@16004: int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq dixon@16004: (* specialised search constructor for assumptions using skips *) dixon@16004: val skip_first_asm_occs_search : dixon@16004: ('a -> 'b -> 'c Seq.seq Seq.seq) -> dixon@16004: 'a -> int -> 'b -> 'c IsaPLib.skipseqT dixon@16004: dixon@16004: (* tactics and methods *) dixon@15936: val eqsubst_asm_meth : int list -> Thm.thm list -> Proof.method dixon@16004: val eqsubst_asm_tac : dixon@16004: int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq dixon@15814: val eqsubst_asm_tac' : dixon@16004: (* search function with skips *) dixon@16004: (searchinfo -> int -> Term.term -> match IsaPLib.skipseqT) dixon@16004: -> int (* skip to *) dixon@16004: -> Thm.thm (* rule *) dixon@16004: -> int (* subgoal number *) dixon@16004: -> Thm.thm (* tgt theorem with subgoals *) dixon@16004: -> Thm.thm Seq.seq (* new theorems *) dixon@15538: dixon@15936: val eqsubst_meth : int list -> Thm.thm list -> Proof.method dixon@16004: val eqsubst_tac : dixon@16004: int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq dixon@15814: val eqsubst_tac' : dixon@16004: (searchinfo -> Term.term -> match Seq.seq) dixon@16004: -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq dixon@15538: dixon@15936: val meth : (bool * int list) * Thm.thm list -> Proof.context -> Proof.method paulson@15481: val setup : (Theory.theory -> Theory.theory) list paulson@15481: end; paulson@15481: dixon@16004: paulson@15481: functor EQSubstTacFUN (structure EqRuleData : EQRULE_DATA) dixon@15538: : EQSUBST_TAC paulson@15481: = struct paulson@15481: dixon@15915: (* a type abriviation for match information *) dixon@15550: type match = dixon@15915: ((Term.indexname * (Term.sort * Term.typ)) list (* type instantiations *) dixon@15915: * (Term.indexname * (Term.typ * Term.term)) list) (* term instantiations *) dixon@15915: * (string * Term.typ) list (* fake named type abs env *) dixon@15915: * (string * Term.typ) list (* type abs env *) dixon@15915: * Term.term (* outer term *) dixon@15550: dixon@16004: type searchinfo = dixon@16004: Sign.sg (* sign for matching *) dixon@16004: * int (* maxidx *) dixon@16004: * BasicIsaFTerm.FcTerm (* focusterm to search under *) dixon@15550: dixon@15538: (* FOR DEBUGGING... dixon@15538: type trace_subst_errT = int (* subgoal *) dixon@15538: * Thm.thm (* thm with all goals *) dixon@15538: * (Thm.cterm list (* certified free var placeholders for vars *) dixon@15538: * Thm.thm) (* trivial thm of goal concl *) dixon@15538: (* possible matches/unifiers *) dixon@15538: * Thm.thm (* rule *) dixon@15538: * (((Term.indexname * Term.typ) list (* type instantiations *) dixon@15538: * (Term.indexname * Term.term) list ) (* term instantiations *) dixon@15538: * (string * Term.typ) list (* Type abs env *) dixon@15538: * Term.term) (* outer term *); dixon@15538: dixon@15538: val trace_subst_err = (ref NONE : trace_subst_errT option ref); dixon@15538: val trace_subst_search = ref false; dixon@15538: exception trace_subst_exp of trace_subst_errT; dixon@15538: *) dixon@15538: dixon@15538: (* also defined in /HOL/Tools/inductive_codegen.ML, dixon@15538: maybe move this to seq.ML ? *) dixon@15538: infix 5 :->; dixon@15538: fun s :-> f = Seq.flat (Seq.map f s); dixon@15538: dixon@15814: (* search from top, left to right, then down *) dixon@15814: fun search_tlr_all_f f ft = paulson@15481: let paulson@15481: fun maux ft = paulson@15481: let val t' = (IsaFTerm.focus_of_fcterm ft) dixon@15538: (* val _ = dixon@15538: if !trace_subst_search then dixon@15538: (writeln ("Examining: " ^ (TermLib.string_of_term t')); dixon@15538: TermLib.writeterm t'; ()) dixon@15538: else (); *) paulson@15481: in paulson@15481: (case t' of dixon@15538: (_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft), dixon@15929: Seq.cons(f ft, paulson@15481: maux (IsaFTerm.focus_right ft))) dixon@15929: | (Abs _) => Seq.cons(f ft, maux (IsaFTerm.focus_abs ft)) dixon@15929: | leaf => Seq.single (f ft)) end paulson@15481: in maux ft end; paulson@15481: dixon@15814: (* search from top, left to right, then down *) dixon@15814: fun search_tlr_valid_f f ft = dixon@15814: let dixon@15814: fun maux ft = dixon@15814: let dixon@15814: val hereseq = if IsaFTerm.valid_match_start ft then f ft else Seq.empty dixon@15814: in dixon@15814: (case (IsaFTerm.focus_of_fcterm ft) of dixon@15814: (_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft), dixon@15929: Seq.cons(hereseq, dixon@15814: maux (IsaFTerm.focus_right ft))) dixon@15929: | (Abs _) => Seq.cons(hereseq, maux (IsaFTerm.focus_abs ft)) dixon@15929: | leaf => Seq.single (hereseq)) dixon@15814: end dixon@15814: in maux ft end; dixon@15814: dixon@15814: (* search all unifications *) dixon@16004: fun searchf_tlr_unify_all (sgn, maxidx, ft) lhs = paulson@15481: IsaFTerm.find_fcterm_matches dixon@15814: search_tlr_all_f dixon@16004: (IsaFTerm.clean_unify_ft sgn maxidx lhs) dixon@16004: ft; paulson@15481: dixon@15814: (* search only for 'valid' unifiers (non abs subterms and non vars) *) dixon@16004: fun searchf_tlr_unify_valid (sgn, maxidx, ft) lhs = dixon@15814: IsaFTerm.find_fcterm_matches dixon@15814: search_tlr_valid_f dixon@16004: (IsaFTerm.clean_unify_ft sgn maxidx lhs) dixon@16004: ft; dixon@15929: 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 *) dixon@15538: fun apply_subst_in_concl i th (cfvs, conclthm) rule m = dixon@15538: (RWInst.rw 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 *) dixon@16004: fun prep_concl_subst i gth = paulson@15481: let paulson@15481: val th = Thm.incr_indexes 1 gth; paulson@15481: val tgt_term = Thm.prop_of th; paulson@15481: paulson@15481: val sgn = Thm.sign_of_thm th; paulson@15481: val ctermify = Thm.cterm_of sgn; paulson@15481: val trivify = Thm.trivial o ctermify; paulson@15481: paulson@15481: val (fixedbody, fvs) = IsaND.fix_alls_term 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@15538: val maxidx = Term.maxidx_of_term conclterm; dixon@16004: val ft = ((IsaFTerm.focus_right dixon@16004: o IsaFTerm.focus_left dixon@16004: o IsaFTerm.fcterm_of_term 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 *) dixon@15814: fun eqsubst_tac' searchf instepthm i th = dixon@15538: let dixon@16004: val (cvfsconclthm, searchinfo) = prep_concl_subst i th; dixon@15538: val stepthms = dixon@15538: Seq.map Drule.zero_var_indexes dixon@15538: (Seq.of_list (EqRuleData.prep_meta_eq instepthm)); dixon@15538: fun rewrite_with_thm r = dixon@15538: let val (lhs,_) = Logic.dest_equals (Thm.concl_of r); dixon@16004: in (searchf searchinfo lhs) dixon@15538: :-> (apply_subst_in_concl i th cvfsconclthm r) end; dixon@16004: in stepthms :-> rewrite_with_thm end; dixon@15538: dixon@15538: dixon@16004: (* Tactic.distinct_subgoals_tac -- fails to free type variables *) dixon@15959: dixon@15959: (* custom version of distinct subgoals that works with term and dixon@15959: type variables. Asssumes th is in beta-eta normal form. dixon@15959: Also, does nothing if flexflex contraints are present. *) dixon@15959: fun distinct_subgoals th = dixon@15959: if List.null (Thm.tpairs_of th) then dixon@15959: let val (fixes,fixedthm) = IsaND.fix_vars_and_tvars th dixon@15959: val (fixedthconcl,cprems) = IsaND.hide_prems fixedthm dixon@15959: in dixon@15959: IsaND.unfix_frees_and_tfrees dixon@15959: fixes dixon@15959: (Drule.implies_intr_list dixon@15959: (Library.gen_distinct dixon@15959: (fn (x, y) => Thm.term_of x = Thm.term_of y) dixon@15959: cprems) fixedthconcl) dixon@15959: end dixon@15959: else th; dixon@15538: dixon@15936: (* General substiuttion of multiple occurances using one of dixon@15936: the given theorems*) dixon@15959: exception eqsubst_occL_exp of dixon@15959: string * (int list) * (Thm.thm list) * int * Thm.thm; dixon@16004: fun skip_first_occs_search occ srchf sinfo lhs = dixon@16004: case (IsaPLib.skipto_seqseq occ (srchf sinfo lhs)) of dixon@16004: IsaPLib.skipmore _ => Seq.empty dixon@16004: | IsaPLib.skipseq ss => Seq.flat ss; dixon@16004: dixon@16004: fun eqsubst_tac occL thms i th = dixon@15936: let val nprems = Thm.nprems_of th in dixon@15936: if nprems < i then Seq.empty else dixon@15936: let val thmseq = (Seq.of_list thms) dixon@15936: fun apply_occ occ th = dixon@15936: thmseq :-> dixon@16004: (fn r => eqsubst_tac' (skip_first_occs_search dixon@15936: occ searchf_tlr_unify_valid) r dixon@15936: (i + ((Thm.nprems_of th) - nprems)) dixon@15936: th); dixon@16004: 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 dixon@15959: end dixon@15959: handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th); 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 *) dixon@15936: fun eqsubst_meth occL inthms = paulson@15481: Method.METHOD dixon@15538: (fn facts => dixon@15936: HEADGOAL ( Method.insert_tac facts THEN' eqsubst_tac occL inthms )); paulson@15481: dixon@16004: (* apply a substitution inside assumption j, keeps asm in the same place *) dixon@16007: fun apply_subst_in_asm i th rule ((cfvs, j, ngoalprems, pth),m) = dixon@16004: let dixon@16004: val th2 = Thm.rotate_rule (j - 1) i th; (* put premice first *) dixon@16004: val preelimrule = dixon@16004: (RWInst.rw m rule pth) dixon@16004: |> (Seq.hd o Tactic.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 *) dixon@16007: (* 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: dixon@16007: (* (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. *) dixon@16004: fun prep_subst_in_asm i gth j = paulson@15481: let paulson@15481: val th = Thm.incr_indexes 1 gth; paulson@15481: val tgt_term = Thm.prop_of th; paulson@15481: paulson@15481: val sgn = Thm.sign_of_thm th; paulson@15481: val ctermify = Thm.cterm_of sgn; paulson@15481: val trivify = Thm.trivial o ctermify; paulson@15481: paulson@15481: val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; paulson@15481: val cfvs = rev (map ctermify fvs); paulson@15481: dixon@15538: val asmt = Library.nth_elem(j - 1,(Logic.strip_imp_prems fixedbody)); dixon@15538: val asm_nprems = length (Logic.strip_imp_prems asmt); dixon@15538: dixon@15538: val pth = trivify asmt; dixon@15538: val maxidx = Term.maxidx_of_term asmt; dixon@15538: dixon@16004: val ft = ((IsaFTerm.focus_right dixon@16004: o IsaFTerm.fcterm_of_term 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 *) dixon@16004: fun prep_subst_in_asms i gth = dixon@16004: map (prep_subst_in_asm i gth) dixon@16004: ((rev o IsaPLib.mk_num_list o length) 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 *) dixon@16004: fun eqsubst_asm_tac' searchf skipocc instepthm i th = dixon@15538: let dixon@16004: val asmpreps = prep_subst_in_asms i th; dixon@15538: val stepthms = dixon@15538: Seq.map Drule.zero_var_indexes dixon@16004: (Seq.of_list (EqRuleData.prep_meta_eq 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) = dixon@16004: (case searchf searchinfo occ lhs of dixon@16004: IsaPLib.skipmore i => occ_search i moreasms dixon@16004: | IsaPLib.skipseq ss => dixon@16004: Seq.append (Seq.map (Library.pair asminfo) dixon@16004: (Seq.flat ss), dixon@16004: occ_search 1 moreasms)) dixon@16004: (* find later substs also *) dixon@16004: in dixon@16004: (occ_search skipocc asmpreps) :-> (apply_subst_in_asm i th r) dixon@16004: end; dixon@16004: in stepthms :-> rewrite_with_thm end; dixon@15538: dixon@16004: dixon@16004: fun skip_first_asm_occs_search searchf sinfo occ lhs = dixon@16004: IsaPLib.skipto_seqseq occ (searchf sinfo lhs); dixon@16004: dixon@16004: fun eqsubst_asm_tac occL thms i th = dixon@16004: let val nprems = Thm.nprems_of th dixon@15538: in dixon@16004: if nprems < i then Seq.empty else dixon@16004: let val thmseq = (Seq.of_list thms) dixon@16004: fun apply_occ occK th = dixon@16004: thmseq :-> dixon@16004: (fn r => dixon@16004: eqsubst_asm_tac' (skip_first_asm_occs_search dixon@16004: searchf_tlr_unify_valid) occK r dixon@16004: (i + ((Thm.nprems_of th) - nprems)) dixon@16004: th); dixon@16004: 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 dixon@16004: end dixon@16004: handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th); 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 *) dixon@15936: fun eqsubst_asm_meth occL inthms = paulson@15481: Method.METHOD dixon@15538: (fn facts => dixon@15936: HEADGOAL (Method.insert_tac facts THEN' eqsubst_asm_tac occL inthms )); paulson@15481: paulson@15481: (* combination method that takes a flag (true indicates that subst paulson@15481: should be done to an assumption, false = apply to the conclusion of paulson@15481: the goal) as well as the theorems to use *) dixon@15936: fun meth ((asmflag, occL), inthms) ctxt = dixon@15936: if asmflag then eqsubst_asm_meth occL inthms else eqsubst_meth occL inthms; paulson@15481: paulson@15481: (* syntax for options, given "(asm)" will give back true, without paulson@15481: gives back false *) paulson@15481: val options_syntax = paulson@15481: (Args.parens (Args.$$$ "asm") >> (K true)) || paulson@15481: (Scan.succeed false); dixon@15936: dixon@15929: val ith_syntax = dixon@15936: (Args.parens (Scan.repeat Args.nat)) dixon@15936: || (Scan.succeed [0]); paulson@15481: paulson@15481: (* method syntax, first take options, then theorems *) paulson@15481: fun meth_syntax meth src ctxt = paulson@15481: meth (snd (Method.syntax ((Scan.lift options_syntax) dixon@15929: -- (Scan.lift ith_syntax) paulson@15481: -- Attrib.local_thms) src ctxt)) paulson@15481: ctxt; paulson@15481: paulson@15481: (* setup function for adding method to theory. *) paulson@15481: val setup = paulson@15481: [Method.add_method ("subst", meth_syntax meth, "Substiution with an equation. Use \"(asm)\" option to substitute in an assumption.")]; paulson@15481: paulson@15481: end;