src/HOL/ex/Quickcheck_Examples.thy
changeset 45927 e0305e4f02c9
parent 45765 cb6ddee6a463
child 45942 4dfb1f6bd99b
--- a/src/HOL/ex/Quickcheck_Examples.thy	Tue Dec 20 17:40:15 2011 +0100
+++ b/src/HOL/ex/Quickcheck_Examples.thy	Tue Dec 20 17:40:17 2011 +0100
@@ -6,7 +6,7 @@
 header {* Examples for the 'quickcheck' command *}
 
 theory Quickcheck_Examples
-imports Complex_Main
+imports Complex_Main "~~/src/HOL/Library/Dlist"
 begin
 
 text {*
@@ -342,6 +342,17 @@
 quickcheck[expect = counterexample]
 oops
 
+subsection {* Examples with abstract types *}
+
+lemma
+  "Dlist.length (Dlist.remove x xs) = Dlist.length xs - 1"
+quickcheck[exhaustive]
+oops
+
+lemma
+  "Dlist.length (Dlist.insert x xs) = Dlist.length xs + 1"
+quickcheck[exhaustive]
+oops
 
 subsection {* Examples with Records *}