Theory HOL.Quickcheck_Narrowing
section ‹Counterexample generator performing narrowing-based testing›
theory Quickcheck_Narrowing
imports Quickcheck_Random
keywords "find_unused_assms" :: diag
begin
subsection ‹Counterexample generator›
subsubsection ‹Code generation setup›
setup ‹Code_Target.add_derived_target ("Haskell_Quickcheck", [(Code_Haskell.target, I)])›
code_printing
code_module Typerep ⇀ (Haskell_Quickcheck) ‹
module Typerep(Typerep(..)) where
data Typerep = Typerep String [Typerep]
› for type_constructor typerep constant Typerep.Typerep
| type_constructor typerep ⇀ (Haskell_Quickcheck) "Typerep.Typerep"
| constant Typerep.Typerep ⇀ (Haskell_Quickcheck) "Typerep.Typerep"
code_reserved Haskell_Quickcheck Typerep
code_printing
type_constructor integer ⇀ (Haskell_Quickcheck) "Prelude.Int"
| constant "0::integer" ⇀
(Haskell_Quickcheck) "!(0/ ::/ Prelude.Int)"
setup ‹
let
val target = "Haskell_Quickcheck";
fun print _ = Code_Haskell.print_numeral "Prelude.Int";
in
Numeral.add_code \<^const_name>‹Code_Numeral.Pos› I print target
#> Numeral.add_code \<^const_name>‹Code_Numeral.Neg› (~) print target
end
›
subsubsection ‹Narrowing's deep representation of types and terms›
datatype (plugins only: code extraction) narrowing_type =
Narrowing_sum_of_products "narrowing_type list list"
datatype (plugins only: code extraction) narrowing_term =
Narrowing_variable "integer list" narrowing_type
| Narrowing_constructor integer "narrowing_term list"
datatype (plugins only: code extraction) (dead 'a) narrowing_cons =
Narrowing_cons narrowing_type "(narrowing_term list ⇒ 'a) list"
primrec map_cons :: "('a => 'b) => 'a narrowing_cons => 'b narrowing_cons"
where
"map_cons f (Narrowing_cons ty cs) = Narrowing_cons ty (map (λc. f ∘ c) cs)"
subsubsection ‹From narrowing's deep representation of terms to \<^theory>‹HOL.Code_Evaluation›'s terms›
class partial_term_of = typerep +
fixes partial_term_of :: "'a itself => narrowing_term => Code_Evaluation.term"
lemma partial_term_of_anything: "partial_term_of x nt ≡ t"
by (rule eq_reflection) (cases "partial_term_of x nt", cases t, simp)
subsubsection ‹Auxilary functions for Narrowing›
consts nth :: "'a list => integer => 'a"
code_printing constant nth ⇀ (Haskell_Quickcheck) infixl 9 "!!"
consts error :: "char list => 'a"
code_printing constant error ⇀ (Haskell_Quickcheck) "error"
consts toEnum :: "integer => char"
code_printing constant toEnum ⇀ (Haskell_Quickcheck) "Prelude.toEnum"
consts marker :: "char"
code_printing constant marker ⇀ (Haskell_Quickcheck) "''\\0'"
subsubsection ‹Narrowing's basic operations›
type_synonym 'a narrowing = "integer => 'a narrowing_cons"
definition cons :: "'a => 'a narrowing"
where
"cons a d = (Narrowing_cons (Narrowing_sum_of_products [[]]) [(λ_. a)])"
fun conv :: "(narrowing_term list => 'a) list => narrowing_term => 'a"
where
"conv cs (Narrowing_variable p _) = error (marker # map toEnum p)"
| "conv cs (Narrowing_constructor i xs) = (nth cs i) xs"
fun non_empty :: "narrowing_type => bool"
where
"non_empty (Narrowing_sum_of_products ps) = (¬ (List.null ps))"
definition "apply" :: "('a => 'b) narrowing => 'a narrowing => 'b narrowing"
where
"apply f a d = (if d > 0 then
(case f d of Narrowing_cons (Narrowing_sum_of_products ps) cfs ⇒
case a (d - 1) of Narrowing_cons ta cas ⇒
let
shallow = non_empty ta;
cs = [(λ(x # xs) ⇒ cf xs (conv cas x)). shallow, cf ← cfs]
in Narrowing_cons (Narrowing_sum_of_products [ta # p. shallow, p ← ps]) cs)
else Narrowing_cons (Narrowing_sum_of_products []) [])"
definition sum :: "'a narrowing => 'a narrowing => 'a narrowing"
where
"sum a b d =
(case a d of Narrowing_cons (Narrowing_sum_of_products ssa) ca ⇒
case b d of Narrowing_cons (Narrowing_sum_of_products ssb) cb ⇒
Narrowing_cons (Narrowing_sum_of_products (ssa @ ssb)) (ca @ cb))"
lemma [fundef_cong]:
assumes "a d = a' d" "b d = b' d" "d = d'"
shows "sum a b d = sum a' b' d'"
using assms unfolding sum_def by (auto split: narrowing_cons.split narrowing_type.split)
lemma [fundef_cong]:
assumes "f d = f' d" "(⋀d'. 0 ≤ d' ∧ d' < d ⟹ a d' = a' d')"
assumes "d = d'"
shows "apply f a d = apply f' a' d'"
proof -
note assms
moreover have "0 < d' ⟹ 0 ≤ d' - 1"
by (simp add: less_integer_def less_eq_integer_def)
ultimately show ?thesis
by (auto simp add: apply_def Let_def
split: narrowing_cons.split narrowing_type.split)
qed
subsubsection ‹Narrowing generator type class›
class narrowing =
fixes narrowing :: "integer => 'a narrowing_cons"