src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
author blanchet
Mon, 10 Sep 2012 17:36:02 +0200
changeset 50277 831d4c259f5f
parent 50250 f9fc2b64c599
child 50278 669a820ef213
permissions -rw-r--r--
debug
blanchet@50138
     1
(*  Title:      HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
blanchet@50138
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@50138
     3
    Copyright   2012
blanchet@50138
     4
blanchet@50138
     5
Tactics for the LFP/GFP sugar.
blanchet@50138
     6
*)
blanchet@50138
     7
blanchet@50138
     8
signature BNF_FP_SUGAR_TACTICS =
blanchet@50138
     9
sig
blanchet@50145
    10
  val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
blanchet@50241
    11
  val mk_coiter_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
blanchet@50176
    12
  val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
blanchet@50138
    13
  val mk_fld_iff_unf_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm
blanchet@50138
    14
    -> tactic
blanchet@50142
    15
  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
blanchet@50141
    16
  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
blanchet@50241
    17
  val mk_iter_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
blanchet@50138
    18
end;
blanchet@50138
    19
blanchet@50138
    20
structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
blanchet@50138
    21
struct
blanchet@50138
    22
blanchet@50140
    23
open BNF_Tactics
blanchet@50138
    24
open BNF_Util
blanchet@50138
    25
blanchet@50145
    26
fun mk_case_tac ctxt n k m case_def ctr_def unf_fld =
blanchet@50145
    27
  Local_Defs.unfold_tac ctxt [case_def, ctr_def, unf_fld] THEN
blanchet@50145
    28
  (rtac (BNF_FP_Util.mk_sum_casesN n k RS ssubst) THEN'
blanchet@50145
    29
   REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
blanchet@50145
    30
   rtac refl) 1;
blanchet@50145
    31
blanchet@50176
    32
fun mk_exhaust_tac ctxt n ctr_defs fld_iff_unf sumEN' =
blanchet@50277
    33
print_tac "A1" THEN(*###*)
blanchet@50277
    34
  Local_Defs.unfold_tac ctxt (fld_iff_unf :: ctr_defs) THEN
blanchet@50277
    35
print_tac ("A2: " ^ Display.string_of_thm ctxt sumEN') THEN(*###*)
blanchet@50277
    36
  rtac sumEN' 1 THEN
blanchet@50277
    37
print_tac "A3" THEN(*###*)
blanchet@50140
    38
  Local_Defs.unfold_tac ctxt @{thms all_prod_eq} THEN
blanchet@50176
    39
  EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, REPEAT_DETERM o dtac @{thm meta_spec},
blanchet@50176
    40
    etac @{thm meta_mp}, atac]) (1 upto n)) 1;
blanchet@50140
    41
blanchet@50138
    42
fun mk_fld_iff_unf_tac ctxt cTs cfld cunf fld_unf unf_fld =
blanchet@50138
    43
  (rtac iffI THEN'
blanchet@50138
    44
   EVERY' (map3 (fn cTs => fn cx => fn th =>
blanchet@50138
    45
     dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
blanchet@50138
    46
     SELECT_GOAL (Local_Defs.unfold_tac ctxt [th]) THEN'
blanchet@50138
    47
     atac) [rev cTs, cTs] [cunf, cfld] [unf_fld, fld_unf])) 1;
blanchet@50138
    48
blanchet@50142
    49
fun mk_half_distinct_tac ctxt fld_inject ctr_defs =
blanchet@50142
    50
  Local_Defs.unfold_tac ctxt (fld_inject :: @{thms sum.inject} @ ctr_defs) THEN
blanchet@50142
    51
  rtac @{thm sum.distinct(1)} 1;
blanchet@50142
    52
blanchet@50141
    53
fun mk_inject_tac ctxt ctr_def fld_inject =
blanchet@50150
    54
  Local_Defs.unfold_tac ctxt [ctr_def] THEN rtac (fld_inject RS ssubst) 1 THEN
blanchet@50150
    55
  Local_Defs.unfold_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
blanchet@50141
    56
blanchet@50226
    57
val iter_like_thms =
blanchet@50248
    58
  @{thms case_unit comp_def convol_def id_apply map_pair_def sum.simps(5,6) sum_map.simps
blanchet@50248
    59
      split_conv};
blanchet@50220
    60
blanchet@50244
    61
fun mk_iter_like_tac pre_map_defs map_ids iter_like_defs fld_iter_like ctr_def ctxt =
blanchet@50244
    62
  Local_Defs.unfold_tac ctxt (ctr_def :: fld_iter_like :: iter_like_defs @ pre_map_defs @ map_ids @
blanchet@50247
    63
    iter_like_thms) THEN Local_Defs.unfold_tac ctxt @{thms id_def} THEN rtac refl 1;
blanchet@50220
    64
blanchet@50228
    65
val coiter_like_ss = ss_only @{thms if_True if_False};
blanchet@50248
    66
val coiter_like_thms = @{thms id_apply map_pair_def sum_map.simps prod.cases};
blanchet@50228
    67
blanchet@50244
    68
fun mk_coiter_like_tac coiter_like_defs map_ids fld_unf_coiter_like pre_map_def ctr_def ctxt =
blanchet@50228
    69
  Local_Defs.unfold_tac ctxt (ctr_def :: coiter_like_defs) THEN
blanchet@50241
    70
  subst_tac ctxt [fld_unf_coiter_like] 1 THEN asm_simp_tac coiter_like_ss 1 THEN
blanchet@50247
    71
  Local_Defs.unfold_tac ctxt (pre_map_def :: coiter_like_thms @ map_ids) THEN
blanchet@50250
    72
  Local_Defs.unfold_tac ctxt @{thms id_def} THEN
blanchet@50250
    73
  (rtac refl ORELSE' subst_tac ctxt @{thms unit_eq} THEN' rtac refl) 1;
blanchet@50228
    74
blanchet@50138
    75
end;