src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
changeset 55180 58a0f8726558
parent 55179 ad7a2cfb8cb2
child 55181 93ab44e992ae
     1.1 --- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML	Wed Oct 02 22:54:42 2013 +0200
     1.2 +++ b/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML	Wed Oct 02 22:54:42 2013 +0200
     1.3 @@ -8,12 +8,12 @@
     1.4  signature BNF_FP_REC_SUGAR_TACTICS =
     1.5  sig
     1.6    val mk_primcorec_assumption_tac: Proof.context -> thm list -> tactic
     1.7 -  val mk_primcorec_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
     1.8 -    int list -> thm list -> tactic
     1.9 -  val mk_primcorec_code_of_raw_tac: thm list -> thm -> tactic
    1.10 +  val mk_primcorec_code_of_raw_code_tac: thm list -> thm -> tactic
    1.11    val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
    1.12    val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
    1.13      tactic
    1.14 +  val mk_primcorec_raw_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list ->
    1.15 +    thm list -> int list -> thm list -> tactic
    1.16    val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
    1.17      thm list -> thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
    1.18    val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic
    1.19 @@ -88,7 +88,7 @@
    1.20    unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl);
    1.21  
    1.22  (* TODO: reduce code duplication with selector tactic above *)
    1.23 -fun mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr =
    1.24 +fun mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr =
    1.25    HEADGOAL (REPEAT o (resolve_tac split_connectI ORELSE' split_tac (split_if :: splits))) THEN
    1.26    mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
    1.27    REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt discIs) THEN
    1.28 @@ -101,11 +101,11 @@
    1.29       eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
    1.30       (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))));
    1.31  
    1.32 -fun mk_primcorec_code_of_ctr_tac ctxt distincts discIs splits split_asms ms ctr_thms =
    1.33 -  EVERY (map2 (mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs splits split_asms)
    1.34 +fun mk_primcorec_raw_code_of_ctr_tac ctxt distincts discIs splits split_asms ms ctr_thms =
    1.35 +  EVERY (map2 (mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms)
    1.36      ms ctr_thms);
    1.37  
    1.38 -fun mk_primcorec_code_of_raw_tac splits raw =
    1.39 +fun mk_primcorec_code_of_raw_code_tac splits raw =
    1.40    HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' REPEAT_DETERM o
    1.41      (rtac refl ORELSE'
    1.42       (TRY o rtac sym) THEN' atac ORELSE'