src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
author blanchet
Fri, 30 Aug 2013 14:17:19 +0200
changeset 54466 c31c0c311cf0
parent 54440 ae49b835ca01
child 54614 75a0427df7a8
permissions -rw-r--r--
more canonical naming
blanchet@54440
     1
(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
blanchet@54440
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@54440
     3
    Copyright   2013
blanchet@54440
     4
blanchet@54440
     5
Tactics for recursor and corecursor sugar.
blanchet@54440
     6
*)
blanchet@54440
     7
blanchet@54440
     8
signature BNF_FP_REC_SUGAR_TACTICS =
blanchet@54440
     9
sig
blanchet@54440
    10
  val mk_primcorec_taut_tac: Proof.context -> tactic
blanchet@54440
    11
  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
blanchet@54440
    12
    tactic
blanchet@54440
    13
  val mk_primcorec_dtr_to_ctr_tac: Proof.context -> int -> thm -> thm -> thm list -> tactic
blanchet@54440
    14
  val mk_primcorec_eq_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
blanchet@54440
    15
    thm list -> thm list -> thm list -> tactic
blanchet@54440
    16
  val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic
blanchet@54440
    17
end;
blanchet@54440
    18
blanchet@54440
    19
structure BNF_FP_Rec_Sugar_Tactics : BNF_FP_REC_SUGAR_TACTICS =
blanchet@54440
    20
struct
blanchet@54440
    21
blanchet@54440
    22
open BNF_Util
blanchet@54440
    23
open BNF_Tactics
blanchet@54440
    24
blanchet@54466
    25
fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
blanchet@54440
    26
  unfold_thms_tac ctxt fun_defs THEN
blanchet@54440
    27
  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
blanchet@54466
    28
  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
blanchet@54440
    29
  HEADGOAL (rtac refl);
blanchet@54440
    30
blanchet@54440
    31
fun mk_primcorec_taut_tac ctxt =
blanchet@54440
    32
  HEADGOAL (etac FalseE) ORELSE
blanchet@54440
    33
  unfold_thms_tac ctxt @{thms de_Morgan_conj de_Morgan_disj not_not not_False_eq_True} THEN
blanchet@54440
    34
  HEADGOAL (SOLVE o REPEAT o (atac ORELSE' resolve_tac @{thms disjI1 disjI2 conjI TrueI}));
blanchet@54440
    35
blanchet@54440
    36
fun mk_primcorec_same_case_tac m =
blanchet@54440
    37
  HEADGOAL (if m = 0 then rtac TrueI
blanchet@54440
    38
    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
blanchet@54440
    39
blanchet@54440
    40
fun mk_primcorec_different_case_tac ctxt excl =
blanchet@54440
    41
  unfold_thms_tac ctxt @{thms not_not not_False_eq_True} THEN
blanchet@54440
    42
  HEADGOAL (rtac excl THEN_ALL_NEW SELECT_GOAL (mk_primcorec_taut_tac ctxt));
blanchet@54440
    43
blanchet@54440
    44
fun mk_primcorec_cases_tac ctxt k m exclsss =
blanchet@54440
    45
  let val n = length exclsss in
blanchet@54440
    46
    EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
blanchet@54440
    47
        | [excl] => mk_primcorec_different_case_tac ctxt excl)
blanchet@54440
    48
      (take k (nth exclsss (k - 1))))
blanchet@54440
    49
  end;
blanchet@54440
    50
blanchet@54440
    51
fun mk_primcorec_prelude ctxt defs thm =
blanchet@54440
    52
  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN unfold_thms_tac ctxt @{thms split};
blanchet@54440
    53
blanchet@54440
    54
fun mk_primcorec_disc_tac ctxt defs disc k m exclsss =
blanchet@54440
    55
  mk_primcorec_prelude ctxt defs disc THEN mk_primcorec_cases_tac ctxt k m exclsss;
blanchet@54440
    56
blanchet@54466
    57
fun mk_primcorec_eq_tac ctxt defs sel k m exclsss maps map_idents map_comps =
blanchet@54440
    58
  mk_primcorec_prelude ctxt defs (sel RS trans) THEN mk_primcorec_cases_tac ctxt k m exclsss THEN
blanchet@54440
    59
  unfold_thms_tac ctxt (@{thms if_if_True if_if_False if_True if_False o_def split_def sum.cases} @
blanchet@54466
    60
    maps @ map_comps @ map_idents) THEN HEADGOAL (rtac refl);
blanchet@54440
    61
blanchet@54440
    62
fun mk_primcorec_dtr_to_ctr_tac ctxt m collapse disc sels =
blanchet@54440
    63
  HEADGOAL (rtac (collapse RS sym RS trans) THEN' rtac disc THEN' REPEAT_DETERM_N m o atac) THEN
blanchet@54440
    64
  unfold_thms_tac ctxt sels THEN HEADGOAL (rtac refl);
blanchet@54440
    65
blanchet@54440
    66
end;