src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
author blanchet
Wed, 02 Oct 2013 10:39:01 +0200
changeset 55161 07ab4fd922c2
parent 55155 bd2e127389f2
child 55179 ad7a2cfb8cb2
permissions -rw-r--r--
strengthen tactic w.r.t. let
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@55040
    10
  val mk_primcorec_assumption_tac: Proof.context -> thm list -> tactic
blanchet@55040
    11
  val mk_primcorec_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
blanchet@55058
    12
    int list -> thm list -> tactic
blanchet@55058
    13
  val mk_primcorec_code_of_raw_tac: thm list -> thm list -> thm -> tactic
panny@54859
    14
  val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> 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@55047
    17
  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
blanchet@55047
    18
    thm list -> thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
blanchet@54440
    19
  val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic
blanchet@54440
    20
end;
blanchet@54440
    21
blanchet@54440
    22
structure BNF_FP_Rec_Sugar_Tactics : BNF_FP_REC_SUGAR_TACTICS =
blanchet@54440
    23
struct
blanchet@54440
    24
blanchet@54440
    25
open BNF_Util
blanchet@54440
    26
open BNF_Tactics
blanchet@54440
    27
blanchet@55042
    28
val falseEs = @{thms not_TrueE FalseE};
blanchet@55047
    29
val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict};
blanchet@55039
    30
val split_if = @{thm split_if};
blanchet@55039
    31
val split_if_asm = @{thm split_if_asm};
blanchet@55039
    32
val split_connectI = @{thms allI impI conjI};
blanchet@55039
    33
blanchet@54466
    34
fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
blanchet@54440
    35
  unfold_thms_tac ctxt fun_defs THEN
blanchet@54440
    36
  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
blanchet@54466
    37
  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
blanchet@54440
    38
  HEADGOAL (rtac refl);
blanchet@54440
    39
blanchet@55040
    40
fun mk_primcorec_assumption_tac ctxt discIs =
blanchet@55040
    41
  HEADGOAL (SELECT_GOAL (unfold_thms_tac ctxt
blanchet@55040
    42
      @{thms not_not not_False_eq_True de_Morgan_conj de_Morgan_disj} THEN
blanchet@55063
    43
    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
blanchet@55040
    44
    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
blanchet@55066
    45
    dresolve_tac discIs THEN' atac ORELSE'
blanchet@55066
    46
    etac notE THEN' atac ORELSE'
blanchet@55066
    47
    etac disjE)))));
blanchet@54440
    48
blanchet@54440
    49
fun mk_primcorec_same_case_tac m =
blanchet@54440
    50
  HEADGOAL (if m = 0 then rtac TrueI
blanchet@54440
    51
    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
blanchet@54440
    52
blanchet@54440
    53
fun mk_primcorec_different_case_tac ctxt excl =
blanchet@55059
    54
  unfold_thms_tac ctxt @{thms not_not not_False_eq_True not_True_eq_False} THEN
blanchet@55040
    55
  HEADGOAL (rtac excl THEN_ALL_NEW SELECT_GOAL (mk_primcorec_assumption_tac ctxt []));
blanchet@54440
    56
blanchet@54440
    57
fun mk_primcorec_cases_tac ctxt k m exclsss =
blanchet@54440
    58
  let val n = length exclsss in
blanchet@54440
    59
    EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
blanchet@54440
    60
        | [excl] => mk_primcorec_different_case_tac ctxt excl)
blanchet@54440
    61
      (take k (nth exclsss (k - 1))))
blanchet@54440
    62
  end;
blanchet@54440
    63
blanchet@54440
    64
fun mk_primcorec_prelude ctxt defs thm =
blanchet@55161
    65
  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN
blanchet@55161
    66
  unfold_thms_tac ctxt @{thms Let_def split};
blanchet@54440
    67
blanchet@54843
    68
fun mk_primcorec_disc_tac ctxt defs disc_corec k m exclsss =
blanchet@54843
    69
  mk_primcorec_prelude ctxt defs disc_corec THEN mk_primcorec_cases_tac ctxt k m exclsss;
blanchet@54440
    70
blanchet@55047
    71
fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms maps map_idents map_comps f_sel k m
blanchet@55047
    72
    exclsss =
blanchet@55038
    73
  mk_primcorec_prelude ctxt defs (f_sel RS trans) THEN
blanchet@54830
    74
  mk_primcorec_cases_tac ctxt k m exclsss THEN
blanchet@55155
    75
  unfold_thms_tac ctxt (@{thms id_apply o_def split_def} @ maps @ map_comps @ map_idents) THEN
blanchet@55155
    76
  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' rtac ext ORELSE'
blanchet@55042
    77
    eresolve_tac falseEs ORELSE'
blanchet@55039
    78
    resolve_tac split_connectI ORELSE'
blanchet@55039
    79
    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
blanchet@55039
    80
    Splitter.split_tac (split_if :: splits) ORELSE'
blanchet@55047
    81
    eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
blanchet@55155
    82
    (CHANGED o SELECT_GOAL (unfold_tac @{thms sum.cases} ctxt)) ORELSE'
blanchet@55037
    83
    etac notE THEN' atac));
blanchet@54440
    84
panny@54859
    85
fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse maybe_disc_f sel_fs =
panny@54857
    86
  HEADGOAL (rtac ((if null sel_fs then collapse else collapse RS sym) RS trans) THEN'
panny@54859
    87
    (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
blanchet@54843
    88
  unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl);
blanchet@54440
    89
blanchet@55045
    90
(* TODO: reduce code duplication with selector tactic above *)
blanchet@55058
    91
fun mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr =
blanchet@55042
    92
  HEADGOAL (REPEAT o (resolve_tac split_connectI ORELSE' split_tac (split_if :: splits))) THEN
blanchet@55042
    93
  mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
blanchet@55042
    94
  REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt discIs) THEN
blanchet@55066
    95
  HEADGOAL (SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
blanchet@55041
    96
    (rtac refl ORELSE' atac ORELSE'
blanchet@55041
    97
     resolve_tac split_connectI ORELSE'
blanchet@55041
    98
     Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
blanchet@55041
    99
     Splitter.split_tac (split_if :: splits) ORELSE'
blanchet@55066
   100
     K (mk_primcorec_assumption_tac ctxt discIs) ORELSE'
blanchet@55047
   101
     eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
blanchet@55066
   102
     (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))));
blanchet@55040
   103
blanchet@55058
   104
fun mk_primcorec_code_of_ctr_tac ctxt distincts discIs splits split_asms ms ctr_thms =
blanchet@55058
   105
  EVERY (map2 (mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs splits split_asms)
blanchet@55058
   106
    ms ctr_thms);
blanchet@54830
   107
blanchet@55058
   108
fun mk_primcorec_code_of_raw_tac splits disc_excludes raw =
blanchet@55041
   109
  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' REPEAT_DETERM o
blanchet@55041
   110
    (rtac refl ORELSE'
blanchet@55041
   111
     (TRY o rtac sym) THEN' atac ORELSE'
blanchet@55041
   112
     resolve_tac split_connectI ORELSE'
blanchet@55041
   113
     Splitter.split_tac (split_if :: splits) ORELSE'
blanchet@55041
   114
     etac notE THEN' atac ORELSE'
blanchet@55058
   115
     (TRY o dresolve_tac disc_excludes) THEN' etac notE THEN' atac));
blanchet@55058
   116
blanchet@54440
   117
end;