# HG changeset patch # User blanchet # Date 1411051660 -7200 # Node ID 7b92932ffea54e36f6547767e633be564ac0ce49 # Parent 1b4d31b7bd105aabdf11fc2d5b22cef6f303561b careful with op = in n2m (actually by Dmitriy Traytel) diff -r 1b4d31b7bd10 -r 7b92932ffea5 src/HOL/Tools/BNF/bnf_fp_n2m.ML --- a/src/HOL/Tools/BNF/bnf_fp_n2m.ML Thu Sep 18 16:47:40 2014 +0200 +++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML Thu Sep 18 16:47:40 2014 +0200 @@ -182,12 +182,13 @@ val castAs = map2 (curry HOLogic.mk_comp) absAs fp_repAs; val castBs = map2 (curry HOLogic.mk_comp) absBs fp_repBs; - val rel_eqs = no_refl (map rel_eq_of_bnf fp_or_nesting_bnfs); + val fp_or_nesting_rel_eqs = no_refl (map rel_eq_of_bnf fp_or_nesting_bnfs); + val fp_or_nesting_rel_monos = map rel_mono_of_bnf fp_or_nesting_bnfs; val rel_xtor_co_induct_thm = mk_rel_xtor_co_induct_thm fp (map3 cast castAs castBs pre_rels) pre_phis rels phis xs ys xtors xtor's (mk_rel_xtor_co_induct_tactic fp abs_inverses rel_xtor_co_inducts rel_defs - rel_monos rel_eqs) + rel_monos fp_or_nesting_rel_eqs fp_or_nesting_rel_monos) lthy; val map_id0s = no_refl (map map_id0_of_bnf bnfs); @@ -209,7 +210,8 @@ in cterm_instantiate_pos cts rel_xtor_co_induct_thm |> singleton (Proof_Context.export names_lthy lthy) - |> unfold_thms lthy (@{thms eq_le_Grp_id_iff all_simps(1,2)[symmetric]} @ rel_eqs) + |> unfold_thms lthy (@{thms eq_le_Grp_id_iff all_simps(1,2)[symmetric]} @ + fp_or_nesting_rel_eqs) |> funpow n (fn thm => thm RS spec) |> unfold_thms lthy (@{thm eq_alt} :: map rel_Grp_of_bnf bnfs @ map_id0s) |> unfold_thms lthy (@{thms vimage2p_id vimage2p_comp comp_apply comp_id @@ -224,7 +226,8 @@ val cts = NONE :: map (SOME o certify lthy) (map HOLogic.eq_const As); in cterm_instantiate_pos cts rel_xtor_co_induct_thm - |> unfold_thms lthy (@{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} @ rel_eqs) + |> unfold_thms lthy (@{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} @ + fp_or_nesting_rel_eqs) |> funpow (2 * n) (fn thm => thm RS spec) |> Conv.fconv_rule (Object_Logic.atomize lthy) |> funpow n (fn thm => thm RS mp) diff -r 1b4d31b7bd10 -r 7b92932ffea5 src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML --- a/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML Thu Sep 18 16:47:40 2014 +0200 +++ b/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML Thu Sep 18 16:47:40 2014 +0200 @@ -8,7 +8,7 @@ signature BNF_FP_N2M_TACTICS = sig val mk_rel_xtor_co_induct_tactic: BNF_Util.fp_kind -> thm list -> thm list -> thm list -> - thm list -> thm list -> {prems: thm list, context: Proof.context} -> tactic + thm list -> thm list -> thm list -> {prems: thm list, context: Proof.context} -> tactic end; structure BNF_FP_N2M_Tactics : BNF_FP_N2M_TACTICS = @@ -20,11 +20,13 @@ val vimage2p_unfolds = o_apply :: @{thms vimage2p_def}; -fun mk_rel_xtor_co_induct_tactic fp abs_inverses co_inducts0 rel_defs rel_monos nesting_rel_eqs - {context = ctxt, prems = raw_C_IHs} = +fun mk_rel_xtor_co_induct_tactic fp abs_inverses co_inducts0 rel_defs rel_monos nesting_rel_eqs0 + nesting_rel_monos0 {context = ctxt, prems = raw_C_IHs} = let - val co_inducts = map (unfold_thms ctxt - (vimage2p_unfolds @ @{thms prod.rel_eq sum.rel_eq} @ nesting_rel_eqs)) co_inducts0; + val nesting_rel_eqs = @{thms prod.rel_eq sum.rel_eq} @ nesting_rel_eqs0; + val nesting_rel_monos = map (fn thm => rotate_prems ~1 (thm RS @{thm predicate2D})) + (@{thms prod.rel_mono sum.rel_mono} @ nesting_rel_monos0); + val co_inducts = map (unfold_thms ctxt (vimage2p_unfolds @ nesting_rel_eqs)) co_inducts0; val unfolds = map (fn def => unfold_thms ctxt (id_apply :: vimage2p_unfolds @ abs_inverses @ no_reflexive [def])) rel_defs; val folded_C_IHs = map (fn thm => thm RS @{thm spec2} RS mp) raw_C_IHs; @@ -41,7 +43,9 @@ HEADGOAL (CONJ_WRAP_GEN' (rtac @{thm context_conjI}) (fn thm => rtac thm THEN_ALL_NEW (rotate_tac ~1 THEN' REPEAT_ALL_NEW (FIRST' [eresolve_tac C_IHs, eresolve_tac C_IH_monos, - rtac @{thm order_refl}, atac, resolve_tac co_inducts]))) + SELECT_GOAL (unfold_thms_tac ctxt nesting_rel_eqs) THEN' rtac @{thm order_refl}, + atac, resolve_tac co_inducts, + resolve_tac C_IH_monos THEN' REPEAT_ALL_NEW (eresolve_tac nesting_rel_monos)]))) co_inducts) end;