diff -r ec949d7206bb -r 7553a1bcecb7 src/HOL/Quickcheck_Exhaustive.thy --- a/src/HOL/Quickcheck_Exhaustive.thy Sat Sep 13 18:08:45 2014 +0200 +++ b/src/HOL/Quickcheck_Exhaustive.thy Sun Sep 14 22:59:30 2014 +0200 @@ -574,8 +574,8 @@ where "pos_bound_cps_if b = (if b then pos_bound_cps_single () else pos_bound_cps_empty)" -datatype (dead 'a) unknown = Unknown | Known 'a -datatype (dead 'a) three_valued = Unknown_value | Value 'a | No_value +datatype (plugins only: code) (dead 'a) unknown = Unknown | Known 'a +datatype (plugins only: code) (dead 'a) three_valued = Unknown_value | Value 'a | No_value type_synonym 'a neg_bound_cps = "('a unknown => term list three_valued) => natural => term list three_valued"