# HG changeset patch # User blanchet # Date 1380747282 -7200 # Node ID ad7a2cfb8cb2addd2b75d5f99ac0f11a03eacef1 # Parent 227908156cd2013f6ffbfe8d5acf49b9d6706a1d got rid of needless argument diff -r 227908156cd2 -r ad7a2cfb8cb2 src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML --- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML Wed Oct 02 22:54:42 2013 +0200 +++ b/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML Wed Oct 02 22:54:42 2013 +0200 @@ -10,7 +10,7 @@ val mk_primcorec_assumption_tac: Proof.context -> thm list -> tactic val mk_primcorec_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list -> thm list -> int list -> thm list -> tactic - val mk_primcorec_code_of_raw_tac: thm list -> thm list -> thm -> tactic + val mk_primcorec_code_of_raw_tac: thm list -> thm -> tactic val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list -> tactic @@ -105,13 +105,12 @@ EVERY (map2 (mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs splits split_asms) ms ctr_thms); -fun mk_primcorec_code_of_raw_tac splits disc_excludes raw = +fun mk_primcorec_code_of_raw_tac splits raw = HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' REPEAT_DETERM o (rtac refl ORELSE' (TRY o rtac sym) THEN' atac ORELSE' resolve_tac split_connectI ORELSE' Splitter.split_tac (split_if :: splits) ORELSE' - etac notE THEN' atac ORELSE' - (TRY o dresolve_tac disc_excludes) THEN' etac notE THEN' atac)); + etac notE THEN' atac)); end;