tactic cleanup
authorblanchet
Thu, 26 Sep 2013 15:13:55 +0200
changeset 53921 46fc95abef09
parent 53920 c6de7f20c845
child 53922 6d40f6e69686
tactic cleanup
src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
--- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML	Thu Sep 26 15:13:28 2013 +0200
+++ b/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML	Thu Sep 26 15:13:55 2013 +0200
@@ -9,8 +9,8 @@
 sig
   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 ->
-    thm list -> int list -> thm list -> tactic
-  val mk_primcorec_code_of_raw_tac: thm list -> thm list -> thm -> thm list -> tactic
+    int list -> thm list -> tactic
+  val mk_primcorec_code_of_raw_tac: thm list -> 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
@@ -83,33 +83,31 @@
     (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
   unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl);
 
-(* TODO: do we need "collapses"? *)
 (* TODO: reduce code duplication with selector tactic above *)
-fun mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs collapses splits split_asms m f_ctr =
+fun mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr =
   HEADGOAL (REPEAT o (resolve_tac split_connectI ORELSE' split_tac (split_if :: splits))) THEN
   mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
   REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt discIs) THEN
   HEADGOAL (SELECT_GOAL (HEADGOAL (REPEAT_DETERM o
     (rtac refl ORELSE' atac ORELSE'
-     eresolve_tac (falseEs @ map (fn thm => thm RS trans) collapses) ORELSE'
      resolve_tac split_connectI ORELSE'
      Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
      Splitter.split_tac (split_if :: splits) ORELSE'
      eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
      (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac))));
 
-fun mk_primcorec_code_of_ctr_tac ctxt distincts discIs collapses splits split_asms ms ctr_thms =
-  EVERY (map2 (mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs collapses splits
-    split_asms) ms ctr_thms);
+fun mk_primcorec_code_of_ctr_tac ctxt distincts discIs splits split_asms ms ctr_thms =
+  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 collapses =
+fun mk_primcorec_code_of_raw_tac splits disc_excludes 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'
-     dresolve_tac disc_excludes THEN' etac notE THEN' atac ORELSE'
-     eresolve_tac (maps (fn thm => [thm, thm RS sym]) collapses)));
+     (TRY o dresolve_tac disc_excludes) THEN' etac notE THEN' atac));
+
 
 end;