# HG changeset patch # User blanchet # Date 1380747282 -7200 # Node ID 58a0f8726558b6a8c66172451b00645a61f919fc # Parent ad7a2cfb8cb2addd2b75d5f99ac0f11a03eacef1 renamings diff -r ad7a2cfb8cb2 -r 58a0f8726558 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 @@ -8,12 +8,12 @@ signature BNF_FP_REC_SUGAR_TACTICS = 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 -> - int list -> thm list -> tactic - val mk_primcorec_code_of_raw_tac: thm list -> thm -> tactic + val mk_primcorec_code_of_raw_code_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 + val mk_primcorec_raw_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list -> + thm list -> int list -> thm list -> tactic val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic @@ -88,7 +88,7 @@ unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl); (* TODO: reduce code duplication with selector tactic above *) -fun mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr = +fun mk_primcorec_raw_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 @@ -101,11 +101,11 @@ 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 splits split_asms ms ctr_thms = - EVERY (map2 (mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs splits split_asms) +fun mk_primcorec_raw_code_of_ctr_tac ctxt distincts discIs splits split_asms ms ctr_thms = + EVERY (map2 (mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms) ms ctr_thms); -fun mk_primcorec_code_of_raw_tac splits raw = +fun mk_primcorec_code_of_raw_code_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'