src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
author blanchet
Wed, 18 Sep 2013 18:12:40 +0200
changeset 54836 7d32f33d340d
parent 54834 221ec353bcc5
child 54839 9b034e6b977c
permissions -rw-r--r--
tuned tactics
     1 (*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2013
     4 
     5 Tactics for recursor and corecursor sugar.
     6 *)
     7 
     8 signature BNF_FP_REC_SUGAR_TACTICS =
     9 sig
    10   val mk_primcorec_code_tac: Proof.context -> thm -> thm list -> tactic
    11   val mk_primcorec_code_of_ctr_tac: Proof.context -> int list -> thm list -> tactic
    12   val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm -> thm list -> tactic
    13   val mk_primcorec_ctr_or_sel_tac: Proof.context -> thm list -> thm -> int -> int ->
    14     thm list list list -> thm list -> thm list -> thm list -> tactic
    15   val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
    16     tactic
    17   val mk_primcorec_assumption_tac: Proof.context -> tactic
    18   val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic
    19 end;
    20 
    21 structure BNF_FP_Rec_Sugar_Tactics : BNF_FP_REC_SUGAR_TACTICS =
    22 struct
    23 
    24 open BNF_Util
    25 open BNF_Tactics
    26 
    27 fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
    28   unfold_thms_tac ctxt fun_defs THEN
    29   HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
    30   unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
    31   HEADGOAL (rtac refl);
    32 
    33 fun mk_primcorec_assumption_tac ctxt =
    34   HEADGOAL (blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt));
    35 
    36 fun mk_primcorec_same_case_tac m =
    37   HEADGOAL (if m = 0 then rtac TrueI
    38     else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
    39 
    40 fun mk_primcorec_different_case_tac ctxt excl =
    41   unfold_thms_tac ctxt @{thms not_not not_False_eq_True} THEN
    42   HEADGOAL (rtac excl THEN_ALL_NEW SELECT_GOAL (mk_primcorec_assumption_tac ctxt));
    43 
    44 fun mk_primcorec_cases_tac ctxt k m exclsss =
    45   let val n = length exclsss in
    46     EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
    47         | [excl] => mk_primcorec_different_case_tac ctxt excl)
    48       (take k (nth exclsss (k - 1))))
    49   end;
    50 
    51 fun mk_primcorec_prelude ctxt defs thm =
    52   unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN unfold_thms_tac ctxt @{thms split};
    53 
    54 fun mk_primcorec_disc_tac ctxt defs disc k m exclsss =
    55   mk_primcorec_prelude ctxt defs disc THEN mk_primcorec_cases_tac ctxt k m exclsss;
    56 
    57 fun mk_primcorec_ctr_or_sel_tac ctxt defs eq_thm k m exclsss maps map_idents map_comps =
    58   mk_primcorec_prelude ctxt defs (eq_thm RS trans) THEN
    59   mk_primcorec_cases_tac ctxt k m exclsss THEN
    60   unfold_thms_tac ctxt (@{thms if_if_True if_if_False if_True if_False if_cancel[of _ True]
    61     if_cancel[of _ False] o_def split_def sum.cases} @ maps @ map_comps @ map_idents) THEN
    62   HEADGOAL (rtac refl);
    63 
    64 fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse disc sels =
    65   HEADGOAL (rtac (collapse RS sym RS trans) THEN' rtac disc THEN' REPEAT_DETERM_N m o atac) THEN
    66   unfold_thms_tac ctxt sels THEN HEADGOAL (rtac refl);
    67 
    68 fun mk_primcorec_code_of_ctr_case_tac ctxt m ctr_thm =
    69   HEADGOAL (REPEAT o rtac @{thm eq_ifI}) THEN
    70   mk_primcorec_prelude ctxt [] (ctr_thm RS trans) THEN
    71   REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt) THEN
    72   HEADGOAL (asm_simp_tac (ss_only @{thms if_True if_False not_False_eq_True} ctxt));
    73 
    74 fun mk_primcorec_code_of_ctr_tac ctxt ms ctr_thms =
    75   EVERY (map2 (mk_primcorec_code_of_ctr_case_tac ctxt) ms ctr_thms);
    76 
    77 fun mk_primcorec_code_tac ctxt raw collapses =
    78   HEADGOAL (rtac (raw RS trans)) THEN unfold_thms_tac ctxt @{thms split_ifs} THEN
    79   Method.intros_tac @{thms conjI impI} [] THEN
    80   REPEAT (HEADGOAL (rtac refl ORELSE' (etac notE THEN' atac) ORELSE'
    81     eresolve_tac (maps (fn thm => [thm, thm RS sym]) collapses)));
    82 
    83 end;