src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
author blanchet
Thu, 26 Sep 2013 01:05:07 +0200
changeset 55038 3d93e8b4ae02
parent 55037 527ece7edc51
child 55039 396999552212
permissions -rw-r--r--
tuning
     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_assumption_tac: Proof.context -> tactic
    11   val mk_primcorec_code_tac: thm list -> thm list -> thm -> thm list -> tactic
    12   val mk_primcorec_code_of_ctr_tac: Proof.context -> thm list -> int list -> thm list -> tactic
    13   val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
    14   val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
    15     tactic
    16   val mk_primcorec_sel_tac: Proof.context -> thm list -> thm -> int -> int ->
    17     thm list list list -> thm list -> thm list -> thm list -> thm list -> thm list ->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_corec k m exclsss =
    55   mk_primcorec_prelude ctxt defs disc_corec THEN mk_primcorec_cases_tac ctxt k m exclsss;
    56 
    57 fun mk_primcorec_sel_tac ctxt defs f_sel k m exclsss maps map_idents map_comps splits split_asms =
    58   mk_primcorec_prelude ctxt defs (f_sel RS trans) THEN
    59   mk_primcorec_cases_tac ctxt k m exclsss THEN
    60   unfold_thms_tac ctxt (@{thms o_def split_def sum.cases} @ maps @ map_comps @ map_idents) THEN
    61   HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE'
    62     eresolve_tac @{thms not_TrueE FalseE} ORELSE'
    63     resolve_tac @{thms allI impI conjI} ORELSE'
    64     Splitter.split_asm_tac (@{thm split_if_asm} :: split_asms) ORELSE'
    65     Splitter.split_tac (@{thm split_if} :: splits) ORELSE'
    66     etac notE THEN' atac));
    67 
    68 fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse maybe_disc_f sel_fs =
    69   HEADGOAL (rtac ((if null sel_fs then collapse else collapse RS sym) RS trans) THEN'
    70     (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
    71   unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl);
    72 
    73 fun mk_primcorec_code_of_ctr_case_tac ctxt eq_caseIs m f_ctr =
    74   HEADGOAL (REPEAT o resolve_tac (@{thm eq_ifI} :: eq_caseIs)) THEN
    75   mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
    76   REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt) THEN
    77   (* FIXME: Something like (ss_only @{thms if_True if_False not_False_eq_True simp_thms} ctxt) *)
    78   HEADGOAL (asm_simp_tac ctxt);
    79 
    80 fun mk_primcorec_code_of_ctr_tac ctxt eq_caseIs ms ctr_thms =
    81   EVERY (map2 (mk_primcorec_code_of_ctr_case_tac ctxt eq_caseIs) ms ctr_thms);
    82 
    83 fun mk_primcorec_code_tac eq_caseIs disc_excludes raw collapses =
    84   HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' REPEAT o
    85     (resolve_tac (maps (fn thm => [thm, thm RS sym]) (@{thm eq_ifI} :: eq_caseIs)) ORELSE'
    86      rtac refl ORELSE' etac notE THEN' atac ORELSE'
    87      dresolve_tac disc_excludes THEN' etac notE THEN' atac ORELSE'
    88      eresolve_tac (maps (fn thm => [thm, thm RS sym]) collapses)));
    89 
    90 end;