src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
changeset 55040 27fd72593624
parent 55039 396999552212
child 55041 446076262e92
     1.1 --- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML	Thu Sep 26 01:05:32 2013 +0200
     1.2 +++ b/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML	Thu Sep 26 02:09:52 2013 +0200
     1.3 @@ -7,9 +7,10 @@
     1.4  
     1.5  signature BNF_FP_REC_SUGAR_TACTICS =
     1.6  sig
     1.7 -  val mk_primcorec_assumption_tac: Proof.context -> tactic
     1.8 +  val mk_primcorec_assumption_tac: Proof.context -> thm list -> tactic
     1.9    val mk_primcorec_code_tac: thm list -> thm list -> thm -> thm list -> tactic
    1.10 -  val mk_primcorec_code_of_ctr_tac: Proof.context -> thm list -> int list -> thm list -> tactic
    1.11 +  val mk_primcorec_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
    1.12 +    thm list -> int list -> thm list -> tactic
    1.13    val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
    1.14    val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
    1.15      tactic
    1.16 @@ -34,8 +35,12 @@
    1.17    unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
    1.18    HEADGOAL (rtac refl);
    1.19  
    1.20 -fun mk_primcorec_assumption_tac ctxt =
    1.21 -  HEADGOAL (blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt));
    1.22 +fun mk_primcorec_assumption_tac ctxt discIs =
    1.23 +  HEADGOAL (SELECT_GOAL (unfold_thms_tac ctxt
    1.24 +      @{thms not_not not_False_eq_True de_Morgan_conj de_Morgan_disj} THEN
    1.25 +    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE'
    1.26 +    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
    1.27 +    dresolve_tac discIs THEN' atac)))));
    1.28  
    1.29  fun mk_primcorec_same_case_tac m =
    1.30    HEADGOAL (if m = 0 then rtac TrueI
    1.31 @@ -43,7 +48,7 @@
    1.32  
    1.33  fun mk_primcorec_different_case_tac ctxt excl =
    1.34    unfold_thms_tac ctxt @{thms not_not not_False_eq_True} THEN
    1.35 -  HEADGOAL (rtac excl THEN_ALL_NEW SELECT_GOAL (mk_primcorec_assumption_tac ctxt));
    1.36 +  HEADGOAL (rtac excl THEN_ALL_NEW SELECT_GOAL (mk_primcorec_assumption_tac ctxt []));
    1.37  
    1.38  fun mk_primcorec_cases_tac ctxt k m exclsss =
    1.39    let val n = length exclsss in
    1.40 @@ -74,14 +79,24 @@
    1.41      (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
    1.42    unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl);
    1.43  
    1.44 -fun mk_primcorec_code_of_ctr_case_tac ctxt splits m f_ctr =
    1.45 +fun mk_primcorec_split distincts discIs collapses splits split_asms =
    1.46 +  HEADGOAL (SELECT_GOAL (HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' atac ORELSE'
    1.47 +    eresolve_tac (@{thms not_TrueE FalseE} @ map (fn thm => thm RS trans) collapses) ORELSE'
    1.48 +    resolve_tac split_connectI ORELSE'
    1.49 +    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
    1.50 +    Splitter.split_tac (split_if :: splits) ORELSE'
    1.51 +    eresolve_tac (map (fn thm => thm RS @{thm neq_eq_eq_contradict}) distincts) THEN' atac ORELSE'
    1.52 +    (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac))));
    1.53 +
    1.54 +fun mk_primcorec_code_of_ctr_single_tac ctxt distincts discIs collapses splits split_asms m f_ctr =
    1.55    HEADGOAL (REPEAT o (resolve_tac split_connectI ORELSE' split_tac (split_if :: splits))) THEN
    1.56    mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
    1.57 -  REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt) THEN
    1.58 -  HEADGOAL (rtac refl);
    1.59 +  REPEAT_DETERM_N m (mk_primcorec_assumption_tac ctxt discIs) THEN
    1.60 +  mk_primcorec_split distincts discIs collapses splits split_asms;
    1.61  
    1.62 -fun mk_primcorec_code_of_ctr_tac ctxt eq_caseIs ms ctr_thms =
    1.63 -  EVERY (map2 (mk_primcorec_code_of_ctr_case_tac ctxt eq_caseIs) ms ctr_thms);
    1.64 +fun mk_primcorec_code_of_ctr_tac ctxt distincs discIs collapses splits split_asms ms ctr_thms =
    1.65 +  EVERY (map2 (mk_primcorec_code_of_ctr_single_tac ctxt distincs discIs collapses splits split_asms)
    1.66 +    ms ctr_thms);
    1.67  
    1.68  fun mk_primcorec_code_tac eq_caseIs disc_excludes raw collapses =
    1.69    HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' REPEAT o