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
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@54830
    10
  val mk_primcorec_code_tac: Proof.context -> thm -> thm list -> tactic
blanchet@54836
    11
  val mk_primcorec_code_of_ctr_tac: Proof.context -> int list -> thm list -> tactic
blanchet@54829
    12
  val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm -> thm list -> tactic
blanchet@54829
    13
  val mk_primcorec_ctr_or_sel_tac: Proof.context -> thm list -> thm -> int -> int ->
blanchet@54829
    14
    thm list list list -> thm list -> thm list -> thm list -> tactic
blanchet@54830
    15
  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
blanchet@54830
    16
    tactic
blanchet@54836
    17
  val mk_primcorec_assumption_tac: Proof.context -> tactic
blanchet@54440
    18
  val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic
blanchet@54440
    19
end;
blanchet@54440
    20
blanchet@54440
    21
structure BNF_FP_Rec_Sugar_Tactics : BNF_FP_REC_SUGAR_TACTICS =
blanchet@54440
    22
struct
blanchet@54440
    23
blanchet@54440
    24
open BNF_Util
blanchet@54440
    25
open BNF_Tactics
blanchet@54440
    26
blanchet@54466
    27
fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
blanchet@54440
    28
  unfold_thms_tac ctxt fun_defs THEN
blanchet@54440
    29
  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
blanchet@54466
    30
  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
blanchet@54440
    31
  HEADGOAL (rtac refl);
blanchet@54440
    32
blanchet@54836
    33
fun mk_primcorec_assumption_tac ctxt =
blanchet@54836
    34
  HEADGOAL (blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt));
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@54836
    42
  HEADGOAL (rtac excl THEN_ALL_NEW SELECT_GOAL (mk_primcorec_assumption_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@54829
    57
fun mk_primcorec_ctr_or_sel_tac ctxt defs eq_thm k m exclsss maps map_idents map_comps =
blanchet@54830
    58
  mk_primcorec_prelude ctxt defs (eq_thm RS trans) THEN
blanchet@54830
    59
  mk_primcorec_cases_tac ctxt k m exclsss THEN
blanchet@54830
    60
  unfold_thms_tac ctxt (@{thms if_if_True if_if_False if_True if_False if_cancel[of _ True]
blanchet@54830
    61
    if_cancel[of _ False] o_def split_def sum.cases} @ maps @ map_comps @ map_idents) THEN
blanchet@54830
    62
  HEADGOAL (rtac refl);
blanchet@54440
    63
blanchet@54829
    64
fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse disc sels =
blanchet@54440
    65
  HEADGOAL (rtac (collapse RS sym RS trans) THEN' rtac disc THEN' REPEAT_DETERM_N m o atac) THEN
blanchet@54440
    66
  unfold_thms_tac ctxt sels THEN HEADGOAL (rtac refl);
blanchet@54440
    67
blanchet@54836
    68
fun mk_primcorec_code_of_ctr_case_tac ctxt m ctr_thm =
blanchet@54830
    69
  HEADGOAL (REPEAT o rtac @{thm eq_ifI}) THEN
blanchet@54836
    70
  mk_primcorec_prelude ctxt [] (ctr_thm RS trans) THEN
blanchet@54836
    71
  REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt) THEN
blanchet@54836
    72
  HEADGOAL (asm_simp_tac (ss_only @{thms if_True if_False not_False_eq_True} ctxt));
blanchet@54830
    73
blanchet@54836
    74
fun mk_primcorec_code_of_ctr_tac ctxt ms ctr_thms =
blanchet@54836
    75
  EVERY (map2 (mk_primcorec_code_of_ctr_case_tac ctxt) ms ctr_thms);
blanchet@54830
    76
blanchet@54830
    77
fun mk_primcorec_code_tac ctxt raw collapses =
blanchet@54830
    78
  HEADGOAL (rtac (raw RS trans)) THEN unfold_thms_tac ctxt @{thms split_ifs} THEN
blanchet@54830
    79
  Method.intros_tac @{thms conjI impI} [] THEN
blanchet@54830
    80
  REPEAT (HEADGOAL (rtac refl ORELSE' (etac notE THEN' atac) ORELSE'
blanchet@54830
    81
    eresolve_tac (maps (fn thm => [thm, thm RS sym]) collapses)));
blanchet@54830
    82
blanchet@54440
    83
end;