wenzelm@41883: (* Title: Tools/case_product.ML wenzelm@41883: Author: Lars Noschinski, TU Muenchen noschinl@41826: noschinl@41826: Combines two case rules into a single one. noschinl@41826: noschinl@41826: Assumes that the theorems are of the form noschinl@41826: "[| C1; ...; Cm; A1 ==> P; ...; An ==> P |] ==> P" noschinl@41826: where m is given by the "consumes" attribute of the theorem. noschinl@41826: *) noschinl@41826: noschinl@41826: signature CASE_PRODUCT = wenzelm@41883: sig noschinl@41826: val combine: Proof.context -> thm -> thm -> thm noschinl@41826: val combine_annotated: Proof.context -> thm -> thm -> thm noschinl@41826: val setup: theory -> theory noschinl@41826: end; noschinl@41826: noschinl@41826: structure Case_Product: CASE_PRODUCT = noschinl@41826: struct noschinl@41826: noschinl@41826: (*Instantiates the conclusion of thm2 to the one of thm1.*) noschinl@41826: fun inst_concl thm1 thm2 = noschinl@41826: let noschinl@41826: val cconcl_of = Drule.strip_imp_concl o Thm.cprop_of noschinl@41826: in Thm.instantiate (Thm.match (cconcl_of thm2, cconcl_of thm1)) thm2 end noschinl@41826: noschinl@41826: fun inst_thms thm1 thm2 ctxt = noschinl@41826: let noschinl@41826: val import = yield_singleton (apfst snd oo Variable.import true) noschinl@41826: val (i_thm1, ctxt') = import thm1 ctxt noschinl@41826: val (i_thm2, ctxt'') = import (inst_concl i_thm1 thm2) ctxt' noschinl@41826: in ((i_thm1, i_thm2), ctxt'') end noschinl@41826: noschinl@41826: (* noschinl@41826: Returns list of prems, where loose bounds have been replaced by frees. noschinl@41826: FIXME: Focus noschinl@41826: *) noschinl@41826: fun free_prems t ctxt = noschinl@41826: let noschinl@41826: val bs = Term.strip_all_vars t noschinl@41826: val (names, ctxt') = Variable.variant_fixes (map fst bs) ctxt noschinl@41826: val subst = map Free (names ~~ map snd bs) noschinl@41826: val t' = map (Term.subst_bounds o pair (rev subst)) (Logic.strip_assums_hyp t) noschinl@41826: in ((t', subst), ctxt') end noschinl@41826: noschinl@41826: fun build_concl_prems thm1 thm2 ctxt = noschinl@41826: let noschinl@41826: val concl = Thm.concl_of thm1 noschinl@41826: noschinl@41826: fun is_consumes t = not (Logic.strip_assums_concl t aconv concl) noschinl@41826: val (p_cons1, p_cases1) = chop_while is_consumes (Thm.prems_of thm1) noschinl@41826: val (p_cons2, p_cases2) = chop_while is_consumes (Thm.prems_of thm2) noschinl@41826: noschinl@41826: val p_cases_prod = map (fn p1 => map (fn p2 => noschinl@41826: let noschinl@41826: val (((t1, subst1), (t2, subst2)), _) = ctxt noschinl@41826: |> free_prems p1 ||>> free_prems p2 noschinl@41826: in noschinl@41826: Logic.list_implies (t1 @ t2, concl) noschinl@41826: |> fold_rev Logic.all (subst1 @ subst2) noschinl@41826: end noschinl@41826: ) p_cases2) p_cases1 noschinl@41826: noschinl@41826: val prems = p_cons1 :: p_cons2 :: p_cases_prod noschinl@41826: in noschinl@41826: (concl, prems) noschinl@41826: end noschinl@41826: noschinl@41826: fun case_product_tac prems struc thm1 thm2 = noschinl@41826: let noschinl@41826: val (p_cons1 :: p_cons2 :: premss) = unflat struc prems noschinl@41826: val thm2' = thm2 OF p_cons2 noschinl@41826: in noschinl@41826: (Tactic.rtac (thm1 OF p_cons1) noschinl@41826: THEN' EVERY' (map (fn p => noschinl@41826: Tactic.rtac thm2' wenzelm@42361: THEN' EVERY' (map (Proof_Context.fact_tac o single) p)) premss) noschinl@41826: ) noschinl@41826: end noschinl@41826: noschinl@41826: fun combine ctxt thm1 thm2 = noschinl@41826: let noschinl@41826: val ((i_thm1, i_thm2), ctxt') = inst_thms thm1 thm2 ctxt noschinl@41826: val (concl, prems_rich) = build_concl_prems i_thm1 i_thm2 ctxt' noschinl@41826: in noschinl@41826: Goal.prove ctxt' [] (flat prems_rich) concl (fn {prems, ...} => noschinl@41826: case_product_tac prems prems_rich i_thm1 i_thm2 1) noschinl@41826: |> singleton (Variable.export ctxt' ctxt) noschinl@41826: end; noschinl@41826: noschinl@41826: fun annotation thm1 thm2 = noschinl@41826: let nipkow@44045: val (cases1, cons1) = apfst (map fst) (Rule_Cases.get thm1) nipkow@44045: val (cases2, cons2) = apfst (map fst) (Rule_Cases.get thm2) nipkow@44045: val names = map_product (fn (x,_) => fn (y,_) => x ^ "_" ^y) cases1 cases2 noschinl@41826: in noschinl@41826: Rule_Cases.case_names names o Rule_Cases.consumes (cons1 + cons2) noschinl@41826: end noschinl@41826: noschinl@41826: fun combine_annotated ctxt thm1 thm2 = noschinl@41826: combine ctxt thm1 thm2 noschinl@41826: |> snd o annotation thm1 thm2 o pair (Context.Proof ctxt) noschinl@41826: noschinl@41826: (* Attribute setup *) noschinl@41826: wenzelm@41883: val case_prod_attr = wenzelm@41883: let noschinl@41826: fun combine_list ctxt = fold (fn x => fn y => combine_annotated ctxt y x) noschinl@41826: in noschinl@41826: Attrib.thms >> (fn thms => Thm.rule_attribute (fn ctxt => fn thm => noschinl@41826: combine_list (Context.proof_of ctxt) thms thm)) noschinl@41826: end noschinl@41826: noschinl@41826: val setup = noschinl@41826: Attrib.setup @{binding "case_product"} case_prod_attr "product with other case rules" noschinl@41826: noschinl@41826: end;