src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
author blanchet
Thu, 26 Sep 2013 01:05:06 +0200
changeset 55037 527ece7edc51
parent 55002 cadccda5be03
child 55038 3d93e8b4ae02
permissions -rw-r--r--
made tactic more flexible w.r.t. case expressions and such
     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_ctr_or_sel_tac: Proof.context -> thm list -> thm -> int -> int ->
    15     thm list list list -> thm list -> thm list -> thm list -> thm list -> thm list ->tactic
    16   val mk_primcorec_disc_of_ctr_tac: thm -> thm -> tactic;
    17   val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
    18     tactic
    19   val mk_primcorec_sel_of_ctr_tac: thm -> thm -> tactic;
    20   val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic
    21 end;
    22 
    23 structure BNF_FP_Rec_Sugar_Tactics : BNF_FP_REC_SUGAR_TACTICS =
    24 struct
    25 
    26 open BNF_Util
    27 open BNF_Tactics
    28 
    29 fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
    30   unfold_thms_tac ctxt fun_defs THEN
    31   HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
    32   unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
    33   HEADGOAL (rtac refl);
    34 
    35 fun mk_primcorec_assumption_tac ctxt =
    36   HEADGOAL (blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt));
    37 
    38 fun mk_primcorec_same_case_tac m =
    39   HEADGOAL (if m = 0 then rtac TrueI
    40     else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
    41 
    42 fun mk_primcorec_different_case_tac ctxt excl =
    43   unfold_thms_tac ctxt @{thms not_not not_False_eq_True} THEN
    44   HEADGOAL (rtac excl THEN_ALL_NEW SELECT_GOAL (mk_primcorec_assumption_tac ctxt));
    45 
    46 fun mk_primcorec_cases_tac ctxt k m exclsss =
    47   let val n = length exclsss in
    48     EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
    49         | [excl] => mk_primcorec_different_case_tac ctxt excl)
    50       (take k (nth exclsss (k - 1))))
    51   end;
    52 
    53 fun mk_primcorec_prelude ctxt defs thm =
    54   unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN unfold_thms_tac ctxt @{thms split};
    55 
    56 fun mk_primcorec_disc_tac ctxt defs disc_corec k m exclsss =
    57   mk_primcorec_prelude ctxt defs disc_corec THEN mk_primcorec_cases_tac ctxt k m exclsss;
    58 
    59 fun mk_primcorec_ctr_or_sel_tac ctxt defs f_eq k m exclsss maps map_idents map_comps splits
    60     split_asms =
    61   mk_primcorec_prelude ctxt defs (f_eq RS trans) THEN
    62   mk_primcorec_cases_tac ctxt k m exclsss THEN
    63   unfold_thms_tac ctxt (@{thms o_def split_def sum.cases} @ maps @ map_comps @ map_idents) THEN
    64   HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE'
    65     eresolve_tac @{thms not_TrueE FalseE} ORELSE'
    66     resolve_tac @{thms allI impI conjI} ORELSE'
    67     Splitter.split_asm_tac (@{thm split_if_asm} :: split_asms) ORELSE'
    68     Splitter.split_tac (@{thm split_if} :: splits) ORELSE'
    69     etac notE THEN' atac));
    70 
    71 fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse maybe_disc_f sel_fs =
    72   HEADGOAL (rtac ((if null sel_fs then collapse else collapse RS sym) RS trans) THEN'
    73     (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
    74   unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl);
    75 
    76 fun mk_primcorec_disc_of_ctr_tac discI f_ctr =
    77   HEADGOAL (rtac discI THEN' rtac f_ctr) THEN ALLGOALS atac;
    78 
    79 fun mk_primcorec_sel_of_ctr_tac sel f_ctr =
    80   HEADGOAL (etac (f_ctr RS arg_cong RS trans) THEN' rtac sel);
    81 
    82 fun mk_primcorec_code_of_ctr_case_tac ctxt eq_caseIs m f_ctr =
    83   HEADGOAL (REPEAT o resolve_tac (@{thm eq_ifI} :: eq_caseIs)) THEN
    84   mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
    85   REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt) THEN
    86   (* FIXME: Something like (ss_only @{thms if_True if_False not_False_eq_True simp_thms} ctxt) *)
    87   HEADGOAL (asm_simp_tac ctxt);
    88 
    89 fun mk_primcorec_code_of_ctr_tac ctxt eq_caseIs ms ctr_thms =
    90   EVERY (map2 (mk_primcorec_code_of_ctr_case_tac ctxt eq_caseIs) ms ctr_thms);
    91 
    92 fun mk_primcorec_code_tac eq_caseIs disc_excludes raw collapses =
    93   HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' REPEAT o
    94     (resolve_tac (maps (fn thm => [thm, thm RS sym]) (@{thm eq_ifI} :: eq_caseIs)) ORELSE'
    95      rtac refl ORELSE' etac notE THEN' atac ORELSE'
    96      dresolve_tac disc_excludes THEN' etac notE THEN' atac ORELSE'
    97      eresolve_tac (maps (fn thm => [thm, thm RS sym]) collapses)));
    98 
    99 end;