wenzelm@32174: (* Title: Provers/splitter.ML nipkow@4: Author: Tobias Nipkow nipkow@1030: Copyright 1995 TU Munich nipkow@4: nipkow@4: Generic case-splitter, suitable for most logics. nipkow@13157: Deals with equalities of the form ?P(f args) = ... nipkow@13157: where "f args" must be a first-order term without duplicate variables. clasohm@0: *) clasohm@0: oheimb@5304: signature SPLITTER_DATA = oheimb@5304: sig wenzelm@32177: val thy : theory oheimb@5553: val mk_eq : thm -> thm webertj@20217: val meta_eq_to_iff: thm (* "x == y ==> x = y" *) webertj@20217: val iffD : thm (* "[| P = Q; Q |] ==> P" *) webertj@20217: val disjE : thm (* "[| P | Q; P ==> R; Q ==> R |] ==> R" *) webertj@20217: val conjE : thm (* "[| P & Q; [| P; Q |] ==> R |] ==> R" *) webertj@20217: val exE : thm (* "[| EX x. P x; !!x. P x ==> Q |] ==> Q" *) webertj@20217: val contrapos : thm (* "[| ~ Q; P ==> Q |] ==> ~ P" *) webertj@20217: val contrapos2 : thm (* "[| Q; ~ P ==> ~ Q |] ==> P" *) webertj@20217: val notnotD : thm (* "~ ~ P ==> P" *) oheimb@5304: end oheimb@5304: oheimb@5304: signature SPLITTER = oheimb@5304: sig webertj@20217: (* somewhat more internal functions *) wenzelm@33242: val cmap_of_split_thms: thm list -> (string * (typ * term * thm * typ * int) list) list wenzelm@33242: val split_posns: (string * (typ * term * thm * typ * int) list) list -> wenzelm@33242: theory -> typ list -> term -> (thm * (typ * typ * int list) list * int list * typ * term) list wenzelm@33242: (* first argument is a "cmap", returns a list of "split packs" *) webertj@20217: (* the "real" interface, providing a number of tactics *) oheimb@5304: val split_tac : thm list -> int -> tactic oheimb@5304: val split_inside_tac: thm list -> int -> tactic oheimb@5304: val split_asm_tac : thm list -> int -> tactic wenzelm@51717: val add_split: thm -> Proof.context -> Proof.context wenzelm@51717: val del_split: thm -> Proof.context -> Proof.context wenzelm@18728: val split_add: attribute wenzelm@18728: val split_del: attribute wenzelm@30513: val split_modifiers : Method.modifier parser list oheimb@5304: end; oheimb@5304: wenzelm@32177: functor Splitter(Data: SPLITTER_DATA): SPLITTER = wenzelm@17881: struct oheimb@5304: wenzelm@18545: val Const (const_not, _) $ _ = wenzelm@35625: Object_Logic.drop_judgment Data.thy wenzelm@18545: (#1 (Logic.dest_implies (Thm.prop_of Data.notnotD))); oheimb@5304: wenzelm@18545: val Const (const_or , _) $ _ $ _ = wenzelm@35625: Object_Logic.drop_judgment Data.thy wenzelm@18545: (#1 (Logic.dest_implies (Thm.prop_of Data.disjE))); wenzelm@18545: wenzelm@35625: val const_Trueprop = Object_Logic.judgment_name Data.thy; wenzelm@18545: berghofe@1721: webertj@20217: fun split_format_err () = error "Wrong format for split rule"; nipkow@4668: oheimb@5553: fun split_thm_info thm = case concl_of (Data.mk_eq thm) of wenzelm@56245: Const(@{const_name Pure.eq}, _) $ (Var _ $ t) $ c => (case strip_comb t of berghofe@13855: (Const p, _) => (p, case c of (Const (s, _) $ _) => s = const_not | _ => false) berghofe@13855: | _ => split_format_err ()) berghofe@13855: | _ => split_format_err (); oheimb@5304: webertj@20217: fun cmap_of_split_thms thms = webertj@20217: let webertj@20217: val splits = map Data.mk_eq thms wenzelm@33242: fun add_thm thm cmap = wenzelm@33242: (case concl_of thm of _ $ (t as _ $ lhs) $ _ => wenzelm@33242: (case strip_comb lhs of (Const(a,aT),args) => wenzelm@33242: let val info = (aT,lhs,thm,fastype_of t,length args) wenzelm@33242: in case AList.lookup (op =) cmap a of wenzelm@33242: SOME infos => AList.update (op =) (a, info::infos) cmap wenzelm@33242: | NONE => (a,[info])::cmap wenzelm@33242: end wenzelm@33242: | _ => split_format_err()) wenzelm@33242: | _ => split_format_err()) webertj@20217: in wenzelm@33242: fold add_thm splits [] webertj@20217: end; webertj@20217: berghofe@54216: val abss = fold (Term.abs o pair ""); berghofe@54216: webertj@20217: (* ------------------------------------------------------------------------- *) webertj@20217: (* mk_case_split_tac *) webertj@20217: (* ------------------------------------------------------------------------- *) webertj@20217: oheimb@5304: fun mk_case_split_tac order = clasohm@0: let clasohm@0: berghofe@1686: (************************************************************ berghofe@1686: Create lift-theorem "trlift" : berghofe@1686: berghofe@7672: [| !!x. Q x == R x; P(%x. R x) == C |] ==> P (%x. Q x) == C berghofe@1686: berghofe@1686: *************************************************************) oheimb@5304: webertj@20217: val meta_iffD = Data.meta_eq_to_iff RS Data.iffD; (* (P == Q) ==> Q ==> P *) webertj@20217: haftmann@22838: val lift = Goal.prove_global Pure.thy ["P", "Q", "R"] wenzelm@24707: [Syntax.read_prop_global Pure.thy "!!x :: 'b. Q(x) == R(x) :: 'c"] wenzelm@24707: (Syntax.read_prop_global Pure.thy "P(%x. Q(x)) == P(%x. R(x))") wenzelm@58838: (fn {context = ctxt, prems} => rewrite_goals_tac ctxt prems THEN resolve_tac [reflexive_thm] 1) nipkow@4: berghofe@54216: val _ $ _ $ (_ $ (_ $ abs_lift) $ _) = prop_of lift; berghofe@54216: clasohm@0: val trlift = lift RS transitive_thm; clasohm@0: clasohm@0: wenzelm@17881: (************************************************************************ berghofe@1686: Set up term for instantiation of P in the lift-theorem wenzelm@17881: berghofe@1686: t : lefthand side of meta-equality in subgoal berghofe@1686: the lift theorem is applied to (see select) berghofe@1686: pos : "path" leading to abstraction, coded as a list berghofe@1686: T : type of body of P(...) berghofe@1686: *************************************************************************) berghofe@1686: berghofe@54216: fun mk_cntxt t pos T = berghofe@54216: let berghofe@54216: fun down [] t = (Bound 0, t) berghofe@54216: | down (p :: ps) t = berghofe@54216: let berghofe@54216: val (h, ts) = strip_comb t berghofe@54216: val (ts1, u :: ts2) = chop p ts berghofe@54216: val (u1, u2) = down ps u berghofe@54216: in berghofe@54216: (list_comb (incr_boundvars 1 h, berghofe@54216: map (incr_boundvars 1) ts1 @ u1 :: berghofe@54216: map (incr_boundvars 1) ts2), berghofe@54216: u2) berghofe@54216: end; berghofe@54216: val (u1, u2) = down (rev pos) t berghofe@54216: in (Abs ("", T, u1), u2) end; nipkow@1030: berghofe@1686: wenzelm@17881: (************************************************************************ berghofe@1686: Set up term for instantiation of P in the split-theorem berghofe@1686: P(...) == rhs berghofe@1686: berghofe@1686: t : lefthand side of meta-equality in subgoal berghofe@1686: the split theorem is applied to (see select) berghofe@1686: T : type of body of P(...) berghofe@4232: tt : the term Const(key,..) $ ... berghofe@1686: *************************************************************************) berghofe@1686: berghofe@4232: fun mk_cntxt_splitthm t tt T = berghofe@4232: let fun repl lev t = wenzelm@52131: if Envir.aeconv(incr_boundvars lev tt, t) then Bound lev berghofe@4232: else case t of berghofe@4232: (Abs (v, T2, t)) => Abs (v, T2, repl (lev+1) t) berghofe@4232: | (Bound i) => Bound (if i>=lev then i+1 else i) berghofe@4232: | (t1 $ t2) => (repl lev t1) $ (repl lev t2) berghofe@4232: | t => t berghofe@4232: in Abs("", T, repl 0 t) end; berghofe@1686: berghofe@1686: berghofe@1686: (* add all loose bound variables in t to list is *) wenzelm@33245: fun add_lbnos t is = add_loose_bnos (t, 0, is); nipkow@1030: berghofe@7672: (* check if the innermost abstraction that needs to be removed nipkow@1064: has a body of type T; otherwise the expansion thm will fail later on nipkow@1064: *) wenzelm@33029: fun type_test (T, lbnos, apsns) = wenzelm@42364: let val (_, U: typ, _) = nth apsns (foldl1 Int.min lbnos) wenzelm@33029: in T = U end; clasohm@0: berghofe@1686: (************************************************************************* berghofe@1686: Create a "split_pack". berghofe@1686: berghofe@1686: thm : the relevant split-theorem, i.e. P(...) == rhs , where P(...) berghofe@1686: is of the form berghofe@1686: P( Const(key,...) $ t_1 $ ... $ t_n ) (e.g. key = "if") berghofe@1686: T : type of P(...) berghofe@7672: T' : type of term to be scanned berghofe@1686: n : number of arguments expected by Const(key,...) berghofe@1686: ts : list of arguments actually found berghofe@1686: apsns : list of tuples of the form (T,U,pos), one tuple for each wenzelm@17881: abstraction that is encountered on the way to the position where berghofe@1686: Const(key, ...) $ ... occurs, where berghofe@1686: T : type of the variable bound by the abstraction berghofe@1686: U : type of the abstraction's body berghofe@1686: pos : "path" leading to the body of the abstraction berghofe@1686: pos : "path" leading to the position where Const(key, ...) $ ... occurs. berghofe@1686: TB : type of Const(key,...) $ t_1 $ ... $ t_n berghofe@1721: t : the term Const(key,...) $ t_1 $ ... $ t_n berghofe@1686: berghofe@1686: A split pack is a tuple of the form berghofe@7672: (thm, apsns, pos, TB, tt) berghofe@1686: Note : apsns is reversed, so that the outermost quantifier's position berghofe@1686: comes first ! If the terms in ts don't contain variables bound berghofe@1686: by other than meta-quantifiers, apsns is empty, because no further berghofe@1686: lifting is required before applying the split-theorem. wenzelm@17881: ******************************************************************************) berghofe@1686: wenzelm@20664: fun mk_split_pack (thm, T: typ, T', n, ts, apsns, pos, TB, t) = nipkow@1064: if n > length ts then [] nipkow@1064: else let val lev = length apsns haftmann@33955: val lbnos = fold add_lbnos (take n ts) [] wenzelm@33317: val flbnos = filter (fn i => i < lev) lbnos berghofe@4232: val tt = incr_boundvars (~lev) t berghofe@7672: in if null flbnos then berghofe@7672: if T = T' then [(thm,[],pos,TB,tt)] else [] berghofe@7672: else if type_test(T,flbnos,apsns) then [(thm, rev apsns,pos,TB,tt)] paulson@2143: else [] nipkow@1064: end; clasohm@0: berghofe@1686: berghofe@1686: (**************************************************************************** blanchet@58318: Recursively scans term for occurrences of Const(key,...) $ ... blanchet@58318: Returns a list of "split-packs" (one for each occurrence of Const(key,...) ) berghofe@1686: berghofe@1686: cmap : association list of split-theorems that should be tried. berghofe@1686: The elements have the format (key,(thm,T,n)) , where berghofe@1686: key : the theorem's key constant ( Const(key,...) $ ... ) berghofe@1686: thm : the theorem itself berghofe@1686: T : type of P( Const(key,...) $ ... ) berghofe@1686: n : number of arguments expected by Const(key,...) berghofe@1686: Ts : types of parameters berghofe@1686: t : the term to be scanned berghofe@1686: ******************************************************************************) berghofe@1686: nipkow@13157: (* Simplified first-order matching; nipkow@13157: assumes that all Vars in the pattern are distinct; nipkow@13157: see Pure/pattern.ML for the full version; nipkow@13157: *) nipkow@13157: local webertj@20217: exception MATCH nipkow@13157: in wenzelm@42367: fun typ_match thy (tyenv, TU) = Sign.typ_match thy TU tyenv wenzelm@42367: handle Type.TYPE_MATCH => raise MATCH; wenzelm@33242: wenzelm@42367: fun fomatch thy args = webertj@20217: let webertj@20217: fun mtch tyinsts = fn webertj@20217: (Ts, Var(_,T), t) => wenzelm@42367: typ_match thy (tyinsts, (T, fastype_of1(Ts,t))) webertj@20217: | (_, Free (a,T), Free (b,U)) => wenzelm@42367: if a=b then typ_match thy (tyinsts,(T,U)) else raise MATCH webertj@20217: | (_, Const (a,T), Const (b,U)) => wenzelm@42367: if a=b then typ_match thy (tyinsts,(T,U)) else raise MATCH webertj@20217: | (_, Bound i, Bound j) => webertj@20217: if i=j then tyinsts else raise MATCH webertj@20217: | (Ts, Abs(_,T,t), Abs(_,U,u)) => wenzelm@42367: mtch (typ_match thy (tyinsts,(T,U))) (U::Ts,t,u) webertj@20217: | (Ts, f$t, g$u) => webertj@20217: mtch (mtch tyinsts (Ts,f,g)) (Ts, t, u) webertj@20217: | _ => raise MATCH webertj@20217: in (mtch Vartab.empty args; true) handle MATCH => false end; wenzelm@33242: end; nipkow@13157: wenzelm@42367: fun split_posns (cmap : (string * (typ * term * thm * typ * int) list) list) thy Ts t = nipkow@6130: let berghofe@7672: val T' = fastype_of1 (Ts, t); berghofe@7672: fun posns Ts pos apsns (Abs (_, T, t)) = berghofe@7672: let val U = fastype_of1 (T::Ts,t) berghofe@7672: in posns (T::Ts) (0::pos) ((T, U, pos)::apsns) t end nipkow@6130: | posns Ts pos apsns t = nipkow@6130: let berghofe@7672: val (h, ts) = strip_comb t wenzelm@33245: fun iter t (i, a) = (i+1, (posns Ts (i::pos) apsns t) @ a); wenzelm@33245: val a = wenzelm@33245: case h of wenzelm@33245: Const(c, cT) => wenzelm@33245: let fun find [] = [] wenzelm@33245: | find ((gcT, pat, thm, T, n)::tups) = wenzelm@42367: let val t2 = list_comb (h, take n ts) in wenzelm@42367: if Sign.typ_instance thy (cT, gcT) andalso fomatch thy (Ts, pat, t2) wenzelm@42367: then mk_split_pack(thm,T,T',n,ts,apsns,pos,type_of1(Ts,t2),t2) wenzelm@42367: else find tups wenzelm@33245: end wenzelm@33245: in find (these (AList.lookup (op =) cmap c)) end wenzelm@33245: | _ => [] wenzelm@33245: in snd (fold iter ts (0, a)) end nipkow@1030: in posns Ts [] [] t end; clasohm@0: webertj@20217: fun shorter ((_,ps,pos,_,_), (_,qs,qos,_,_)) = wenzelm@4519: prod_ord (int_ord o pairself length) (order o pairself length) wenzelm@4519: ((ps, pos), (qs, qos)); wenzelm@4519: berghofe@1686: berghofe@1686: (************************************************************ berghofe@1686: call split_posns with appropriate parameters berghofe@1686: *************************************************************) clasohm@0: nipkow@1030: fun select cmap state i = wenzelm@42367: let wenzelm@42367: val thy = Thm.theory_of_thm state wenzelm@42367: val goal = term_of (Thm.cprem_of state i); wenzelm@42367: val Ts = rev (map #2 (Logic.strip_params goal)); wenzelm@42367: val _ $ t $ _ = Logic.strip_assums_concl goal; wenzelm@42367: in (Ts, t, sort shorter (split_posns cmap thy Ts t)) end; nipkow@1030: wenzelm@42367: fun exported_split_posns cmap thy Ts t = wenzelm@42367: sort shorter (split_posns cmap thy Ts t); berghofe@1686: berghofe@1686: (************************************************************* berghofe@1686: instantiate lift theorem berghofe@1686: berghofe@1686: if t is of the form berghofe@1686: ... ( Const(...,...) $ Abs( .... ) ) ... berghofe@1686: then berghofe@1686: P = %a. ... ( Const(...,...) $ a ) ... berghofe@1686: where a has type T --> U berghofe@1686: berghofe@1686: Ts : types of parameters berghofe@1686: t : lefthand side of meta-equality in subgoal berghofe@1686: the split theorem is applied to (see cmap) berghofe@1686: T,U,pos : see mk_split_pack berghofe@1686: state : current proof state berghofe@1686: i : no. of subgoal berghofe@1686: **************************************************************) berghofe@1686: berghofe@7672: fun inst_lift Ts t (T, U, pos) state i = berghofe@7672: let wenzelm@22578: val cert = cterm_of (Thm.theory_of_thm state); berghofe@54216: val (cntxt, u) = mk_cntxt t pos (T --> U); berghofe@54216: val trlift' = Thm.lift_rule (Thm.cprem_of state i) berghofe@54216: (Thm.rename_boundvars abs_lift u trlift); berghofe@54216: val (P, _) = strip_comb (fst (Logic.dest_equals berghofe@54216: (Logic.strip_assums_concl (Thm.prop_of trlift')))); berghofe@54216: in cterm_instantiate [(cert P, cert (abss Ts cntxt))] trlift' berghofe@7672: end; clasohm@0: clasohm@0: berghofe@1686: (************************************************************* berghofe@1686: instantiate split theorem berghofe@1686: berghofe@1686: Ts : types of parameters berghofe@1686: t : lefthand side of meta-equality in subgoal berghofe@1686: the split theorem is applied to (see cmap) berghofe@4232: tt : the term Const(key,..) $ ... berghofe@1686: thm : the split theorem berghofe@1686: TB : type of body of P(...) berghofe@1686: state : current proof state berghofe@4232: i : number of subgoal berghofe@1686: **************************************************************) berghofe@1686: berghofe@4232: fun inst_split Ts t tt thm TB state i = wenzelm@17881: let wenzelm@18145: val thm' = Thm.lift_rule (Thm.cprem_of state i) thm; berghofe@7672: val (P, _) = strip_comb (fst (Logic.dest_equals wenzelm@22596: (Logic.strip_assums_concl (Thm.prop_of thm')))); wenzelm@22578: val cert = cterm_of (Thm.theory_of_thm state); berghofe@7672: val cntxt = mk_cntxt_splitthm t tt TB; wenzelm@33245: in cterm_instantiate [(cert P, cert (abss Ts cntxt))] thm' berghofe@4232: end; berghofe@1686: berghofe@7672: berghofe@1686: (***************************************************************************** berghofe@1686: The split-tactic wenzelm@17881: berghofe@1686: splits : list of split-theorems to be tried berghofe@1686: i : number of subgoal the tactic should be applied to berghofe@1686: *****************************************************************************) berghofe@1686: clasohm@0: fun split_tac [] i = no_tac clasohm@0: | split_tac splits i = webertj@20217: let val cmap = cmap_of_split_thms splits berghofe@54216: fun lift_tac Ts t p st = compose_tac (false, inst_lift Ts t p st i, 2) i st berghofe@7672: fun lift_split_tac state = berghofe@7672: let val (Ts, t, splits) = select cmap state i nipkow@1030: in case splits of berghofe@7672: [] => no_tac state berghofe@7672: | (thm, apsns, pos, TB, tt)::_ => nipkow@1030: (case apsns of berghofe@7672: [] => compose_tac (false, inst_split Ts t tt thm TB state i, 0) i state berghofe@7672: | p::_ => EVERY [lift_tac Ts t p, wenzelm@58838: resolve_tac [reflexive_thm] (i+1), berghofe@7672: lift_split_tac] state) nipkow@1030: end wenzelm@17881: in COND (has_fewer_prems i) no_tac wenzelm@58838: (resolve_tac [meta_iffD] i THEN lift_split_tac) clasohm@0: end; clasohm@0: webertj@20217: in (split_tac, exported_split_posns) end; (* mk_case_split_tac *) berghofe@1721: oheimb@5304: wenzelm@33242: val (split_tac, split_posns) = mk_case_split_tac int_ord; oheimb@4189: wenzelm@33242: val (split_inside_tac, _) = mk_case_split_tac (rev_order o int_ord); oheimb@5304: oheimb@4189: oheimb@4189: (***************************************************************************** oheimb@4189: The split-tactic for premises wenzelm@17881: oheimb@4189: splits : list of split-theorems to be tried oheimb@5304: ****************************************************************************) wenzelm@33242: fun split_asm_tac [] = K no_tac wenzelm@17881: | split_asm_tac splits = oheimb@5304: berghofe@13855: let val cname_list = map (fst o fst o split_thm_info) splits; wenzelm@17881: fun tac (t,i) = wenzelm@20664: let val n = find_index (exists_Const (member (op =) cname_list o #1)) wenzelm@17881: (Logic.strip_assums_hyp t); wenzelm@56245: fun first_prem_is_disj (Const (@{const_name Pure.imp}, _) $ (Const (c, _) wenzelm@18545: $ (Const (s, _) $ _ $ _ )) $ _ ) = c = const_Trueprop andalso s = const_or wenzelm@56245: | first_prem_is_disj (Const(@{const_name Pure.all},_)$Abs(_,_,t)) = wenzelm@17881: first_prem_is_disj t wenzelm@17881: | first_prem_is_disj _ = false; webertj@20217: (* does not work properly if the split variable is bound by a quantifier *) wenzelm@17881: fun flat_prems_tac i = SUBGOAL (fn (t,i) => wenzelm@17881: (if first_prem_is_disj t wenzelm@58838: then EVERY[eresolve_tac [Data.disjE] i, rotate_tac ~1 i, wenzelm@17881: rotate_tac ~1 (i+1), wenzelm@17881: flat_prems_tac (i+1)] wenzelm@17881: else all_tac) wenzelm@17881: THEN REPEAT (eresolve_tac [Data.conjE,Data.exE] i) wenzelm@17881: THEN REPEAT (dresolve_tac [Data.notnotD] i)) i; webertj@20217: in if n<0 then no_tac else (DETERM (EVERY' wenzelm@58838: [rotate_tac n, eresolve_tac [Data.contrapos2], wenzelm@17881: split_tac splits, wenzelm@58838: rotate_tac ~1, eresolve_tac [Data.contrapos], rotate_tac ~1, webertj@20217: flat_prems_tac] i)) wenzelm@17881: end; oheimb@4189: in SUBGOAL tac oheimb@4189: end; oheimb@4189: nipkow@10652: fun gen_split_tac [] = K no_tac nipkow@10652: | gen_split_tac (split::splits) = nipkow@10652: let val (_,asm) = split_thm_info split nipkow@10652: in (if asm then split_asm_tac else split_tac) [split] ORELSE' nipkow@10652: gen_split_tac splits nipkow@10652: end; wenzelm@8468: wenzelm@18688: wenzelm@8468: (** declare split rules **) wenzelm@8468: wenzelm@45620: (* add_split / del_split *) wenzelm@8468: wenzelm@33242: fun string_of_typ (Type (s, Ts)) = wenzelm@33242: (if null Ts then "" else enclose "(" ")" (commas (map string_of_typ Ts))) ^ s berghofe@13859: | string_of_typ _ = "_"; berghofe@13859: wenzelm@17881: fun split_name (name, T) asm = "split " ^ berghofe@13859: (if asm then "asm " else "") ^ name ^ " :: " ^ string_of_typ T; oheimb@4189: wenzelm@51717: fun add_split split ctxt = wenzelm@33242: let wenzelm@45620: val (name, asm) = split_thm_info split wenzelm@45620: val tac = (if asm then split_asm_tac else split_tac) [split] wenzelm@52037: in Simplifier.addloop (ctxt, (split_name name asm, K tac)) end; berghofe@1721: wenzelm@51717: fun del_split split ctxt = wenzelm@45620: let val (name, asm) = split_thm_info split wenzelm@51717: in Simplifier.delloop (ctxt, split_name name asm) end; berghofe@1721: wenzelm@8468: wenzelm@8468: (* attributes *) wenzelm@8468: wenzelm@8468: val splitN = "split"; wenzelm@8468: wenzelm@45620: val split_add = Simplifier.attrib add_split; wenzelm@45620: val split_del = Simplifier.attrib del_split; wenzelm@8634: wenzelm@58826: val _ = wenzelm@58826: Theory.setup wenzelm@58826: (Attrib.setup @{binding split} wenzelm@58826: (Attrib.add_del split_add split_del) "declare case split rule"); wenzelm@58826: wenzelm@8634: wenzelm@9703: (* methods *) wenzelm@8468: wenzelm@8468: val split_modifiers = wenzelm@58048: [Args.$$$ splitN -- Args.colon >> K (Method.modifier split_add @{here}), wenzelm@58048: Args.$$$ splitN -- Args.add -- Args.colon >> K (Method.modifier split_add @{here}), wenzelm@58048: Args.$$$ splitN -- Args.del -- Args.colon >> K (Method.modifier split_del @{here})]; wenzelm@8468: wenzelm@58826: val _ = wenzelm@58826: Theory.setup wenzelm@58826: (Method.setup @{binding split} wenzelm@58826: (Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (CHANGED_PROP o gen_split_tac ths)))) wenzelm@58826: "apply case split rule"); oheimb@4189: berghofe@1721: end;