Theory Quickcheck_Examples
section ‹Examples for the 'quickcheck' command›
theory Quickcheck_Examples
imports Complex_Main "HOL-Library.Dlist" "HOL-Library.DAList_Multiset"
begin
text ‹
The 'quickcheck' command allows to find counterexamples by evaluating
formulae.
Currently, there are two different exploration schemes:
- random testing: this is incomplete, but explores the search space faster.
- exhaustive testing: this is complete, but increasing the depth leads to
exponentially many assignments.
quickcheck can handle quantifiers on finite universes.
›
declare [[quickcheck_timeout = 3600]]
subsection ‹Lists›
theorem "map g (map f xs) = map (g o f) xs"
quickcheck[random, expect = no_counterexample]
quickcheck[exhaustive, size = 3, expect = no_counterexample]
oops
theorem "map g (map f xs) = map (f o g) xs"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
oops
theorem "rev (xs @ ys) = rev ys @ rev xs"
quickcheck[random, expect = no_counterexample]
quickcheck[exhaustive, expect = no_counterexample]
quickcheck[exhaustive, size = 1000, timeout = 0.1]
oops
theorem "rev (xs @ ys) = rev xs @ rev ys"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
oops
theorem "rev (rev xs) = xs"
quickcheck[random, expect = no_counterexample]
quickcheck[exhaustive, expect = no_counterexample]
oops
theorem "rev xs = xs"
quickcheck[tester = random, finite_types = true, report = false, expect = counterexample]
quickcheck[tester = random, finite_types = false, report = false, expect = counterexample]
quickcheck[tester = random, finite_types = true, report = true, expect = counterexample]
quickcheck[tester = random, finite_types = false, report = true, expect = counterexample]
quickcheck[tester = exhaustive, finite_types = true, expect = counterexample]
quickcheck[tester = exhaustive, finite_types = false, expect = counterexample]
oops
text ‹An example involving functions inside other data structures›
primrec app :: "('a ⇒ 'a) list ⇒ 'a ⇒ 'a" where
"app [] x = x"
| "app (f # fs) x = app fs (f x)"
lemma "app (fs @ gs) x = app gs (app fs x)"
quickcheck[random, expect = no_counterexample]
quickcheck[exhaustive, size = 2, expect = no_counterexample]
by (induct fs arbitrary: x) simp_all
lemma "app (fs @ gs) x = app fs (app gs x)"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
oops
primrec occurs :: "'a ⇒ 'a list ⇒ nat" where
"occurs a [] = 0"
| "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
primrec del1 :: "'a ⇒ 'a list ⇒ 'a list" where
"del1 a [] = []"
| "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))"
text ‹A lemma, you'd think to be true from our experience with delAll›
lemma "Suc (occurs a (del1 a xs)) = occurs a xs"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
oops
lemma "xs ~= [] ⟶ Suc (occurs a (del1 a xs)) = occurs a xs"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
oops
lemma "0 < occurs a xs ⟶ Suc (occurs a (del1 a xs)) = occurs a xs"
quickcheck[random, expect = no_counterexample]
quickcheck[exhaustive, expect = no_counterexample]
by (induct xs) auto
primrec replace :: "'a ⇒ 'a ⇒ 'a list ⇒ 'a list" where
"replace a b [] = []"
| "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs))
else (x#(replace a b xs)))"
lemma "occurs a xs = occurs b (replace a b xs)"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
oops
lemma "occurs b xs = 0 ∨ a=b ⟶ occurs a xs = occurs b (replace a b xs)"
quickcheck[random, expect = no_counterexample]
quickcheck[exhaustive, expect = no_counterexample]
by (induct xs) simp_all
subsection ‹Trees›