src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML
changeset 55990 41c6b99c5fb7
parent 55642 63beb38e9258
child 57279 88263522c31e
equal deleted inserted replaced
55989:55827fc7c0dd 55990:41c6b99c5fb7
   138     eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
   138     eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
   139     etac notE THEN' atac ORELSE'
   139     etac notE THEN' atac ORELSE'
   140     (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def comp_def split_def
   140     (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def comp_def split_def
   141        sum.case} @ mapsx @ map_comps @ map_idents))) ORELSE'
   141        sum.case} @ mapsx @ map_comps @ map_idents))) ORELSE'
   142     fo_rtac @{thm cong} ctxt ORELSE'
   142     fo_rtac @{thm cong} ctxt ORELSE'
   143     rtac ext));
   143     rtac @{thm ext}));
   144 
   144 
   145 fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs =
   145 fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs =
   146   HEADGOAL (rtac ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN'
   146   HEADGOAL (rtac ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN'
   147     (the_default (K all_tac) (Option.map rtac disc_fun_opt)) THEN' REPEAT_DETERM_N m o atac) THEN
   147     (the_default (K all_tac) (Option.map rtac disc_fun_opt)) THEN' REPEAT_DETERM_N m o atac) THEN
   148   unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac refl);
   148   unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac refl);