added theory to provide workaround to support nested datatypes in quickcheck (until quickcheck is generalized to support it with new datatypes)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Quickcheck_Examples/Quickcheck_Nesting.thy Thu Sep 01 12:10:52 2016 +0200
@@ -0,0 +1,19 @@
+theory Quickcheck_Nesting
+imports Main
+begin
+
+ML \<open>
+let
+ open BNF_FP_Def_Sugar
+ open BNF_LFP_Compat
+
+ val compat_plugin = Plugin_Name.declare_setup @{binding compat};
+
+ fun compat fp_sugars =
+ perhaps (try (datatype_compat (map (fst o dest_Type o #T) fp_sugars)));
+in
+ Theory.setup (fp_sugars_interpretation compat_plugin compat)
+end
+\<close>
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Quickcheck_Examples/Quickcheck_Nesting_Example.thy Thu Sep 01 12:10:52 2016 +0200
@@ -0,0 +1,11 @@
+theory Quickcheck_Nesting_Example
+imports Quickcheck_Nesting
+begin
+
+datatype x = X "x list"
+
+lemma "X a = X b"
+quickcheck[exhaustive, size = 4, expect = counterexample]
+oops
+
+end
--- a/src/HOL/ROOT Thu Sep 01 11:41:10 2016 +0200
+++ b/src/HOL/ROOT Thu Sep 01 12:10:52 2016 +0200
@@ -978,6 +978,7 @@
Quickcheck_Lattice_Examples
Completeness
Quickcheck_Interfaces
+ Quickcheck_Nesting_Example
theories [condition = ISABELLE_GHC]
Hotel_Example
Quickcheck_Narrowing_Examples