# HG changeset patch # User blanchet # Date 1347909210 -7200 # Node ID 6d05b8452cf31ae001f3f1b1c1bce0bf36382f5c # Parent f27f83f71e94e6298ff612ce15b2241486cd3000 got rid of one "auto" in induction tactic diff -r f27f83f71e94 -r 6d05b8452cf3 src/HOL/Codatatype/BNF_FP.thy --- a/src/HOL/Codatatype/BNF_FP.thy Mon Sep 17 16:57:22 2012 +0200 +++ b/src/HOL/Codatatype/BNF_FP.thy Mon Sep 17 21:13:30 2012 +0200 @@ -104,11 +104,25 @@ "sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)" by simp -lemma UN_compreh_bex: -"\{y. \x \ A. y = {}} = {}" -"\{y. \x \ A. y = {f x}} = {y. \x \ A. y = f x}" +lemma UN_compreh_bex_eq_empty: +"\{y. \x\A. y = {}} = {}" +by blast + +lemma UN_compreh_bex_eq_singleton: +"\{y. \x\A. y = {f x}} = {y. \x\A. y = f x}" +by blast + +lemma mem_UN_comprehI: +"z \ {y. \x\A. y = f x} \ z \ \{y. \x\A. y = {f x}}" +"z \ {y. \x\A. y = f x} \ B \ z \ \{y. \x\A. y = {f x}} \ B" +"z \ \{y. \x\A. y = F x} \ \{y. \x\A. y = G x} \ z \ \{y. \x\A. y = F x \ G x}" +"z \ \{y. \x\A. y = F x} \ (\{y. \x\A. y = G x} \ B) \ z \ \{y. \x\A. y = F x \ G x} \ B" by blast+ +lemma mem_UN_comprehI': +"\y. (\x\A. y = F x) \ z \ y \ z \ \{y. \x\A. y = {y. \y'\F x. y = y'}}" +by blast + lemma induct_set_step: "\B \ A; c \ f B\ \ \C. (\a \ A. C = f a) \ c \ C" apply (rule exI) apply (rule conjI) diff -r f27f83f71e94 -r 6d05b8452cf3 src/HOL/Codatatype/Tools/bnf_fp_sugar.ML --- a/src/HOL/Codatatype/Tools/bnf_fp_sugar.ML Mon Sep 17 16:57:22 2012 +0200 +++ b/src/HOL/Codatatype/Tools/bnf_fp_sugar.ML Mon Sep 17 21:13:30 2012 +0200 @@ -499,7 +499,9 @@ val nesting_map_ids = map map_id_of_bnf nesting_bnfs; fun mk_map live Ts Us t = - let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in + let + val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last + in Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t end; diff -r f27f83f71e94 -r 6d05b8452cf3 src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML --- a/src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML Mon Sep 17 16:57:22 2012 +0200 +++ b/src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML Mon Sep 17 21:13:30 2012 +0200 @@ -97,10 +97,17 @@ etac @{thm induct_set_step}) THEN' atac ORELSE' SELECT_GOAL (auto_tac ctxt); +fun gen_UN_comprehI_tac UnI = REPEAT_DETERM o resolve_tac @{thms mem_UN_comprehI} THEN' rtac UnI; + +fun gen_UnIN_tac 1 1 = REPEAT_DETERM o resolve_tac @{thms mem_UN_comprehI'} + | gen_UnIN_tac _ 1 = gen_UN_comprehI_tac @{thm UnI1} THEN' gen_UnIN_tac 1 1 + | gen_UnIN_tac n k = gen_UN_comprehI_tac @{thm UnI2} THEN' gen_UnIN_tac (n - 1) (k - 1); + val induct_prem_prem_thms = - @{thms SUP_empty Sup_empty Sup_insert UN_compreh_bex UN_insert Un_assoc[symmetric] Un_empty_left - Un_empty_right Union_Un_distrib collect_def[abs_def] fst_conv image_def o_apply snd_conv - snd_prod_fun sum.cases sup_bot_right fst_map_pair map_pair_simp sum_map.simps}; + @{thms SUP_empty Sup_empty Sup_insert UN_compreh_bex_eq_empty UN_compreh_bex_eq_singleton + UN_insert Un_assoc Un_empty_left Un_empty_right Union_Un_distrib collect_def[abs_def] fst_conv + image_def o_apply snd_conv snd_prod_fun sum.cases sup_bot_right fst_map_pair map_pair_simp + sum_map.simps}; (* These rules interfere with the "set_natural'" properties of "sum" and "prod", so we explicitly delay them. *) @@ -114,8 +121,7 @@ (pre_set_defs @ set_natural's @ induct_prem_prem_thms)), SELECT_GOAL (Local_Defs.unfold_tac ctxt (induct_prem_prem_thms_delayed @ induct_prem_prem_thms)), - rtac (mk_UnIN pp jj) THEN' mk_induct_prem_prem_endgame_tac ctxt qq ORELSE' - SELECT_GOAL (auto_tac ctxt)]) (rev ppjjqqkks)) 1; + gen_UnIN_tac pp jj THEN' mk_induct_prem_prem_endgame_tac ctxt qq]) (rev ppjjqqkks)) 1; fun mk_induct_discharge_prem_tac ctxt nn n set_natural's pre_set_defs m k ppjjqqkks = let val r = length ppjjqqkks in