src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
author blanchet
Tue, 04 Sep 2012 16:27:27 +0200
changeset 50142 f7326a0d7f19
parent 50141 1bbd7a37fc29
child 50145 3c26e17b2849
permissions -rw-r--r--
implemented "mk_half_distinct_tac"
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@50140
    10
  val mk_exhaust_tac: Proof.context -> int -> int list -> thm list -> thm -> thm -> tactic
blanchet@50138
    11
  val mk_fld_iff_unf_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm
blanchet@50138
    12
    -> tactic
blanchet@50142
    13
  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
blanchet@50141
    14
  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
blanchet@50138
    15
end;
blanchet@50138
    16
blanchet@50138
    17
structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
blanchet@50138
    18
struct
blanchet@50138
    19
blanchet@50140
    20
open BNF_Tactics
blanchet@50138
    21
open BNF_Util
blanchet@50138
    22
blanchet@50140
    23
fun mk_exhaust_tac ctxt n ms ctr_defs fld_iff_unf sumEN' =
blanchet@50140
    24
  Local_Defs.unfold_tac ctxt (fld_iff_unf :: ctr_defs) THEN
blanchet@50140
    25
  rtac sumEN' 1 THEN
blanchet@50140
    26
  Local_Defs.unfold_tac ctxt @{thms all_prod_eq} THEN
blanchet@50140
    27
  EVERY' (map2 (fn k => fn m =>
blanchet@50140
    28
     select_prem_tac n (REPEAT_DETERM_N m o dtac @{thm meta_spec} THEN' etac @{thm meta_mp}) k THEN'
blanchet@50140
    29
     atac) (1 upto n) ms) 1;
blanchet@50140
    30
blanchet@50138
    31
fun mk_fld_iff_unf_tac ctxt cTs cfld cunf fld_unf unf_fld =
blanchet@50138
    32
  (rtac iffI THEN'
blanchet@50138
    33
   EVERY' (map3 (fn cTs => fn cx => fn th =>
blanchet@50138
    34
     dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
blanchet@50138
    35
     SELECT_GOAL (Local_Defs.unfold_tac ctxt [th]) THEN'
blanchet@50138
    36
     atac) [rev cTs, cTs] [cunf, cfld] [unf_fld, fld_unf])) 1;
blanchet@50138
    37
blanchet@50142
    38
fun mk_half_distinct_tac ctxt fld_inject ctr_defs =
blanchet@50142
    39
  Local_Defs.unfold_tac ctxt (fld_inject :: @{thms sum.inject} @ ctr_defs) THEN
blanchet@50142
    40
  rtac @{thm sum.distinct(1)} 1;
blanchet@50142
    41
blanchet@50141
    42
fun mk_inject_tac ctxt ctr_def fld_inject =
blanchet@50141
    43
  Local_Defs.unfold_tac ctxt [ctr_def] THEN
blanchet@50141
    44
  rtac (fld_inject RS ssubst) 1 THEN
blanchet@50141
    45
  Local_Defs.unfold_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN
blanchet@50141
    46
  rtac refl 1;
blanchet@50141
    47
blanchet@50138
    48
end;