--- a/src/HOL/BNF/BNF_FP.thy Tue Apr 23 17:13:14 2013 +0200
+++ b/src/HOL/BNF/BNF_FP.thy Tue Apr 23 17:15:44 2013 +0200
@@ -116,6 +116,9 @@
"\<Union>{y. \<exists>x\<in>A. y = {x}} = A"
by blast+
+lemma Inl_Inr_False: "(Inl x = Inr y) = False"
+by simp
+
lemma prod_set_simps:
"fsts (x, y) = {x}"
"snds (x, y) = {y}"
--- a/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML Tue Apr 23 17:13:14 2013 +0200
+++ b/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML Tue Apr 23 17:15:44 2013 +0200
@@ -550,8 +550,8 @@
fun mk_rel_thm postproc ctr_defs' cxIn cyIn =
fold_thms lthy ctr_defs'
- (unfold_thms lthy (pre_rel_def :: (if lfp then [] else [dtor_ctor]) @
- sum_prod_thms_rel)
+ (unfold_thms lthy (@{thm Inl_Inr_False} :: pre_rel_def ::
+ (if lfp then [] else [dtor_ctor]) @ sum_prod_thms_rel)
(cterm_instantiate_pos (nones @ [SOME cxIn, SOME cyIn]) fp_rel_thm))
|> postproc
|> singleton (Proof_Context.export names_lthy no_defs_lthy);