--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML Tue Sep 04 13:06:41 2012 +0200
@@ -0,0 +1,26 @@
+(* Title: HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Tactics for the LFP/GFP sugar.
+*)
+
+signature BNF_FP_SUGAR_TACTICS =
+sig
+ val mk_fld_iff_unf_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm
+ -> tactic
+end;
+
+structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
+struct
+
+open BNF_Util
+
+fun mk_fld_iff_unf_tac ctxt cTs cfld cunf fld_unf unf_fld =
+ (rtac iffI THEN'
+ EVERY' (map3 (fn cTs => fn cx => fn th =>
+ dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
+ SELECT_GOAL (Local_Defs.unfold_tac ctxt [th]) THEN'
+ atac) [rev cTs, cTs] [cunf, cfld] [unf_fld, fld_unf])) 1;
+
+end;