testable
Class
Testable propositions have a procedure that can generate counter-examples together with a proof that they invalidate the proposition.
This is a port of the Haskell QuickCheck library.
Creating Customized Instances
The type classes testable
and sampleable
are the means by which
slim_check
creates samples and tests them. For instance, the proposition
∀ i j : ℕ, i ≤ j
has a testable
instance because ℕ
is sampleable
and i ≤ j
is decidable. Once slim_check
finds the testable
instance, it can start using the instance to repeatedly creating samples
and checking whether they satisfy the property. This allows the
user to create new instances and apply slim_check
to new situations.
Polymorphism
The property testable.check (∀ (α : Type) (xs ys : list α), xs ++ ys
= ys ++ xs)
shows us that type-polymorphic properties can be
tested. α
is instantiated with ℤ
first and then tested as normal
monomorphic properties.
The monomorphisation limits the applicability of slim_check
to
polymorphic properties that can be stated about integers. The
limitation may be lifted in the future but, for now, if
one wishes to use a different type than ℤ
, one has to refer to
the desired type.
What do I do if I'm testing a property about my newly defined type?
Let us consider a type made for a new formalization:
structure my_type :=
(x y : ℕ)
(h : x ≤ y)
How do we test a property about my_type
? For instance, let us consider
testable.check $ ∀ a b : my_type, a.y ≤ b.x → a.x ≤ b.y
. Writing this
property as is will give us an error because we do not have an instance
of sampleable my_type
. We can define one as follows:
instance : sampleable my_type :=
{ sample := do
x ← sample ℕ,
xy_diff ← sample ℕ,
return { x := x, y := x + xy_diff, h := /- some proof -/ } }
We can see that the instance is very simple because our type is built
up from other type that have sampleable
instances. sampleable
also
has a shrink
method but it is optional. We may want to implement one
for ease of testing as:
/- ... -/
shrink := λ ⟨x,y,h⟩, (λ ⟨x,y⟩, { x := x, y := x + y, h := /- proof -/}) <$> shrink (x, y - x)
}
Again, we take advantage of the fact that other types have useful
shrink
implementations, in this case prod
.
Optimizing the sampling
Some properties are guarded by a proposition. For instance, recall this example:
#eval testable.check (∀ x : ℕ, 2 ∣ x → x < 100)
When testing the above example, we generate a natural number, we check
that it is even and test it if it is even or throw it away and start
over otherwise. Statistically, we can expect half of our samples to be
thrown away by such a filter. Sometimes, the filter is more
restrictive. For instance we might need x
to be a prime
number. This would cause most of our samples to be discarded.
We can help slim_check
find good samples by providing specialized
sampleable instances. Below, we show an instance for the subtype
of even natural numbers. This means that, when producing
a sample, it is forced to produce a proof that it is even.
instance {k : ℕ} [fact (0 < k)] : sampleable { x : ℕ // k ∣ x } :=
{ sample := do { n ← sample ℕ, pure ⟨k*n, dvd_mul_right _ _⟩ },
shrink := λ ⟨x,h⟩, (λ y, ⟨k*y, dvd_mul_right _ _⟩) <$> shrink x }
Such instance will be preferred when testing a proposition of the shape
∀ x : T, p x → q
We can observe the effect by enabling tracing:
/- no specialized sampling -/
#eval testable.check (∀ x : ℕ, 2 ∣ x → x < 100) { enable_tracing := tt }
-- discard
-- x := 1
-- discard
-- x := 41
-- discard
-- x := 3
-- discard
-- x := 5
-- discard
-- x := 5
-- discard
-- x := 197
-- discard
-- x := 469
-- discard
-- x := 9
-- discard
-- ===================
-- Found problems!
-- x := 552
-- -------------------
/- let us define a specialized sampling instance -/
instance {k : ℕ} : sampleable { x : ℕ // k ∣ x } :=
{ sample := do { n ← sample ℕ, pure ⟨k*n, dvd_mul_right _ _⟩ },
shrink := λ ⟨x,h⟩, (λ y, ⟨k*y, dvd_mul_right _ _⟩) <$> shrink x }
#eval testable.check (∀ x : ℕ, 2 ∣ x → x < 100) { enable_tracing := tt }
-- ===================
-- Found problems!
-- x := 358
-- -------------------
Similarly, it is common to write properties of the form: ∀ i j, i ≤ j → ...
as the following example show:
#eval check (∀ i j k : ℕ, j < k → i - k < i - j)
Without subtype instances, the above property discards many samples
because j < k
does not hold. Fortunately, we have appropriate
instance to choose k
intelligently.
Main definitions
testable
classtestable.check
: a way to test a proposition using random examples
Tags
random testing
References
- https://hackage.haskell.org/package/QuickCheck
- success : Π (p : Prop), psum unit p → slim_check.test_result p
- gave_up : Π (p : Prop), ℕ → slim_check.test_result p
- failure : Π (p : Prop), ¬p → list string → slim_check.test_result p
Result of trying to disprove p
The constructors are:
success : (psum unit p) → test_result
succeed when we find another example satisfyingp
Insuccess h
,h
is an optional proof of the proposition. Without the proof, all we know is that we found one example wherep
holds. With a proof, the one test was sufficient to prove thatp
holds and we do not need to keep finding examples.gave_up {} : ℕ → test_result
give up when a well-formed example cannot be generated.gave_up n
tells us thatn
invalid examples were tried. Above 100, we give up on the proposition and report that we did not find a way to properly test it.failure : ¬ p → (list string) → test_result
a counter-example top
; the strings specify values for the relevant variables.failure h vs
also carries a proof thatp
does not hold. This way, we can guarantee no false positive.
format a test_result
as a string.
Equations
- (slim_check.test_result.failure a vs).to_string = "failed " ++ to_string vs ++ ""
- (slim_check.test_result.gave_up n).to_string = "gave up " ++ to_string n ++ " times"
- (slim_check.test_result.success (psum.inr h)).to_string = "success (with proof)"
- (slim_check.test_result.success (psum.inl ())).to_string = "success (without proof)"
Equations
- run : bool → bool → slim_check.gen (slim_check.test_result p)
testable p
uses random examples to try to disprove p
.
Instances
- slim_check.decidable_testable
- slim_check.and_testable
- slim_check.or_testable
- slim_check.iff_testable
- slim_check.imp_dec_testable
- slim_check.var_testable
- slim_check.eq_testable
- slim_check.all_types_testable
- slim_check.exists_testable
- slim_check.subtype_var_testable
- slim_check.unused_var_testable
- slim_check.test_forall_in_list
- slim_check.exists_testable'
applicative combinator proof carrying test results
Combine the test result for properties p
and q
to create a test for their conjunction.
Equations
- slim_check.and_counter_example (slim_check.test_result.failure Hce xs) (slim_check.test_result.failure a a_1) = slim_check.test_result.failure _ xs
- slim_check.and_counter_example (slim_check.test_result.failure Hce xs) (slim_check.test_result.gave_up a) = slim_check.test_result.failure _ xs
- slim_check.and_counter_example (slim_check.test_result.failure Hce xs) (slim_check.test_result.success a) = slim_check.test_result.failure _ xs
- slim_check.and_counter_example (slim_check.test_result.gave_up a) (slim_check.test_result.failure Hce xs) = slim_check.test_result.failure _ xs
- slim_check.and_counter_example (slim_check.test_result.gave_up n) (slim_check.test_result.gave_up m) = slim_check.test_result.gave_up (n + m)
- slim_check.and_counter_example (slim_check.test_result.gave_up n) (slim_check.test_result.success a) = slim_check.test_result.gave_up n
- slim_check.and_counter_example (slim_check.test_result.success a) (slim_check.test_result.failure Hce xs) = slim_check.test_result.failure _ xs
- slim_check.and_counter_example (slim_check.test_result.success a) (slim_check.test_result.gave_up n) = slim_check.test_result.gave_up n
- slim_check.and_counter_example (slim_check.test_result.success xs) (slim_check.test_result.success ys) = slim_check.test_result.success (slim_check.combine (slim_check.combine (psum.inr and.intro) xs) ys)
Combine the test result for properties p
and q
to create a test for their disjunction
Equations
- slim_check.or_counter_example (slim_check.test_result.failure Hce xs) (slim_check.test_result.failure Hce' ys) = slim_check.test_result.failure _ (xs ++ ys)
- slim_check.or_counter_example (slim_check.test_result.failure a a_1) (slim_check.test_result.gave_up n) = slim_check.test_result.gave_up n
- slim_check.or_counter_example (slim_check.test_result.failure a a_1) (slim_check.test_result.success ys) = slim_check.test_result.success (slim_check.combine (psum.inr or.inr) ys)
- slim_check.or_counter_example (slim_check.test_result.gave_up n) (slim_check.test_result.failure a a_1) = slim_check.test_result.gave_up n
- slim_check.or_counter_example (slim_check.test_result.gave_up n) (slim_check.test_result.gave_up m) = slim_check.test_result.gave_up (n + m)
- slim_check.or_counter_example (slim_check.test_result.gave_up a) (slim_check.test_result.success ys) = slim_check.test_result.success (slim_check.combine (psum.inr or.inr) ys)
- slim_check.or_counter_example (slim_check.test_result.success xs) (slim_check.test_result.failure a a_1) = slim_check.test_result.success (slim_check.combine (psum.inr or.inl) xs)
- slim_check.or_counter_example (slim_check.test_result.success xs) (slim_check.test_result.gave_up a) = slim_check.test_result.success (slim_check.combine (psum.inr or.inl) xs)
- slim_check.or_counter_example (slim_check.test_result.success xs) (slim_check.test_result.success a) = slim_check.test_result.success (slim_check.combine (psum.inr or.inl) xs)
If q → p
, then ¬ p → ¬ q
which means that testing p
can allow us
to find counter-examples to q
.
Equations
- slim_check.convert_counter_example h (slim_check.test_result.failure Hce xs) _x = slim_check.test_result.failure _ xs
- slim_check.convert_counter_example h (slim_check.test_result.gave_up n) _x = slim_check.test_result.gave_up n
- slim_check.convert_counter_example h (slim_check.test_result.success Hp) Hpq = slim_check.test_result.success (slim_check.combine Hpq Hp)
Test q
by testing p
and proving the equivalence between the two.
Equations
When we assign a value to a universally quantified variable, we record that value using this function so that our counter-examples can be informative.
Equations
- slim_check.add_to_counter_example x h (slim_check.test_result.failure Hce xs) _x = slim_check.test_result.failure _ (x :: xs)
- slim_check.add_to_counter_example x h (slim_check.test_result.gave_up a) hpq = slim_check.convert_counter_example h (slim_check.test_result.gave_up a) hpq
- slim_check.add_to_counter_example x h (slim_check.test_result.success a) hpq = slim_check.convert_counter_example h (slim_check.test_result.success a) hpq
Add some formatting to the information recorded by add_to_counter_example
.
Equations
- slim_check.add_var_to_counter_example var x h = slim_check.add_to_counter_example (var ++ " := " ++ to_string x) h
Gadget used to introspect the name of bound variables.
It is used with the testable
typeclass so that
testable (named_binder "x" (∀ x, p x))
can use the variable name
of x
in error messages displayed to the user. If we find that instantiating
the above quantifier with 3 falsifies it, we can print:
==============
Problem found!
==============
x := 3
Equations
- slim_check.named_binder n p = p
Is the given test result a failure?
Equations
- slim_check.and_testable p q = {run := λ (tracing min : bool), slim_check.testable.run p tracing min >>= λ (xp : slim_check.test_result p), slim_check.testable.run q tracing min >>= λ (xq : slim_check.test_result q), pure (slim_check.and_counter_example xp xq)}
Equations
- slim_check.or_testable p q = {run := λ (tracing min : bool), slim_check.testable.run p tracing min >>= λ (xp : slim_check.test_result p), slim_check.or_testable._match_1 p q tracing min xp xp}
- slim_check.or_testable._match_1 p q tracing min xp (slim_check.test_result.failure a a_1) = slim_check.testable.run q tracing min >>= λ (xq : slim_check.test_result q), pure (slim_check.or_counter_example xp xq)
- slim_check.or_testable._match_1 p q tracing min xp (slim_check.test_result.gave_up a) = slim_check.testable.run q tracing min >>= λ (xq : slim_check.test_result q), pure (slim_check.or_counter_example xp xq)
- slim_check.or_testable._match_1 p q tracing min xp (slim_check.test_result.success (psum.inr h)) = pure (slim_check.test_result.success (psum.inr _))
- slim_check.or_testable._match_1 p q tracing min xp (slim_check.test_result.success (psum.inl h)) = pure (slim_check.test_result.success (psum.inl h))
Equations
- slim_check.iff_testable p q = {run := λ (tracing min : bool), slim_check.testable.run (p ∧ q ∨ ¬p ∧ ¬q) tracing min >>= λ (xp : slim_check.test_result (p ∧ q ∨ ¬p ∧ ¬q)), return (slim_check.convert_counter_example' _ xp)}
Equations
- slim_check.imp_dec_testable var p β = {run := λ (tracing min : bool), dite p (λ (h : p), (λ (r : slim_check.test_result (β h)), slim_check.convert_counter_example _ r (psum.inr _)) <$> slim_check.testable.run (β h) tracing min) (λ (h : ¬p), ite ↥tracing (trace "discard" (λ («_» : unit), return (slim_check.test_result.gave_up 1))) (return (slim_check.test_result.gave_up 1)))}
Equations
- slim_check.all_types_testable var f = {run := λ (tracing min : bool), slim_check.testable.run (f ℤ) tracing min >>= λ (r : slim_check.test_result (f ℤ)), return (slim_check.add_var_to_counter_example var "ℤ" _ r)}
Trace the value of sampled variables if the sample is discarded.
Equations
- slim_check.trace_if_giveup tracing_enabled var val (slim_check.test_result.failure a a_1) = λ (_x : thunk β), _x ()
- slim_check.trace_if_giveup tracing_enabled var val (slim_check.test_result.gave_up _x) = ite ↥tracing_enabled (trace (" " ++ to_string var ++ (" := " ++ to_string val ++ ""))) (λ (_x : thunk β), _x ())
- slim_check.trace_if_giveup tracing_enabled var val (slim_check.test_result.success a) = λ (_x : thunk β), _x ()
testable instance for a property iterating over the element of a list
Equations
- slim_check.test_forall_in_list var var' α β (x :: xs) = {run := λ (tracing min : bool), slim_check.testable.run (β x) tracing min >>= λ (r : slim_check.test_result (β x)), slim_check.trace_if_giveup tracing var x r (λ («_» : unit), slim_check.test_forall_in_list._match_1 var var' α β x xs tracing min r (slim_check.test_forall_in_list var var' α β xs) (slim_check.test_forall_in_list var var' α β xs) r)}
- slim_check.test_forall_in_list var var' α β list.nil = {run := λ (tracing min : bool), return (slim_check.test_result.success (psum.inr _))}
- slim_check.test_forall_in_list._match_1 var var' α β x xs tracing min r _f_1 _f_2 (slim_check.test_result.failure _x _x_1) = return (slim_check.add_var_to_counter_example var x _ r)
- slim_check.test_forall_in_list._match_1 var var' α β x xs tracing min r _f_1 _f_2 (slim_check.test_result.gave_up n) = slim_check.testable.run (slim_check.named_binder var (∀ (x : α), slim_check.named_binder var' (x ∈ xs → β x))) tracing min >>= λ (rs : slim_check.test_result (slim_check.named_binder var (∀ (x : α), slim_check.named_binder var' (x ∈ xs → β x)))), slim_check.test_forall_in_list._match_2 var var' α β x xs n rs
- slim_check.test_forall_in_list._match_1 var var' α β x xs tracing min r _f_1 _f_2 (slim_check.test_result.success hp) = slim_check.testable.run (slim_check.named_binder var (∀ (x : α), slim_check.named_binder var' (x ∈ xs → β x))) tracing min >>= λ (rs : slim_check.test_result (slim_check.named_binder var (∀ (x : α), slim_check.named_binder var' (x ∈ xs → β x)))), return (slim_check.convert_counter_example _ rs (slim_check.combine (psum.inr _) hp))
- slim_check.test_forall_in_list._match_2 var var' α β x xs n (slim_check.test_result.failure Hce xs_1) = return (slim_check.test_result.failure _ xs_1)
- slim_check.test_forall_in_list._match_2 var var' α β x xs n (slim_check.test_result.gave_up n') = return (slim_check.test_result.gave_up (n + n'))
- slim_check.test_forall_in_list._match_2 var var' α β x xs n (slim_check.test_result.success _x) = return (slim_check.test_result.gave_up n)
Test proposition p
by randomly selecting one of the provided
testable instances.
Equations
- slim_check.combine_testable p t h = {run := λ (tracing min : bool), have this : 0 < (list.map (λ (t : slim_check.testable p), slim_check.testable.run p tracing min) t).length, from _, slim_check.gen.one_of (list.map (λ (t : slim_check.testable p), slim_check.testable.run p tracing min) t) this}
Shrink a counter-example x
by using shrink x
, picking the first
candidate that falsifies a property and recursively shrinking that one.
The process is guaranteed to terminate because shrink x
produces
a proof that all the values it produces are smaller (according to sizeof
)
than x
.
Equations
- slim_check.minimize_aux α β = _.fix (λ (x : α) (f_rec : Π (y : α), has_well_founded.r y x → option_t slim_check.gen (Σ (x : α), slim_check.test_result (β x))), lazy_list.mfirst (λ (_x : {y // sizeof y < sizeof x}), slim_check.minimize_aux._match_2 α β x _x) (slim_check.sampleable.shrink x) >>= λ (_p : Σ (a : α), slim_check.test_result (β a) × plift (slim_check.sizeof_lt a x)), slim_check.minimize_aux._match_3 α β x f_rec _p)
- slim_check.minimize_aux._match_2 α β x ⟨a, h⟩ = monad_lift (uliftable.up (slim_check.testable.run (β a) ff tt)) >>= λ (_p : ulift (slim_check.test_result (β a))), slim_check.minimize_aux._match_1 α β x a h _p
- slim_check.minimize_aux._match_3 α β x f_rec ⟨y, (h₀, {down := h₁})⟩ = (f_rec y h₁ <|> pure ⟨y, h₀⟩)
- slim_check.minimize_aux._match_1 α β x a h {down := r} = ite ↥(slim_check.is_failure r) (pure ⟨a, (r, {down := h})⟩) failure
Once a property fails to hold on an example, look for smaller counter-examples to show the user.
Equations
- slim_check.minimize α β x r = (slim_check.minimize_aux α (λ (x : α), β x) x).run >>= λ (x' : option (Σ (x : α), slim_check.test_result (β x))), pure (x'.get_or_else ⟨x, r⟩)
Equations
- slim_check.exists_testable var var' α β p = {run := λ (tracing min : bool), slim_check.testable.run (slim_check.named_binder var (∀ (x : α), slim_check.named_binder var' (β x → p))) tracing min >>= λ (x : slim_check.test_result (slim_check.named_binder var (∀ (x : α), slim_check.named_binder var' (β x → p)))), pure (slim_check.convert_counter_example' _ x)}
Equations
- slim_check.exists_testable' var var' α β xs p = {run := λ (tracing min : bool), slim_check.testable.run (slim_check.named_binder var (∀ (x : α), slim_check.named_binder "H" (x ∈ xs → slim_check.named_binder var' (β x → p)))) tracing min >>= λ (x : slim_check.test_result (slim_check.named_binder var (∀ (x : α), slim_check.named_binder "H" (x ∈ xs → slim_check.named_binder var' (β x → p))))), pure (slim_check.convert_counter_example' _ x)}
Test a universal property by creating a sample of the right type and instantiating the bound variable with it.
Equations
- slim_check.var_testable var α β = {run := λ (tracing min : bool), uliftable.adapt_down (slim_check.sampleable.sample α) (λ (x : α), slim_check.testable.run (β x) tracing ff >>= λ (r : slim_check.test_result (β x)), uliftable.adapt_down (ite (↥(slim_check.is_failure r) ∧ ↥min) (slim_check.minimize α (λ (x : α), β x) x r) (pure ⟨x, r⟩)) (λ (_x : Σ (x : α), slim_check.test_result (β x)), slim_check.var_testable._match_1 var α β tracing _x))}
- slim_check.var_testable._match_1 var α β tracing ⟨x, r⟩ = return (slim_check.trace_if_giveup tracing var x r (λ («_» : unit), (slim_check.add_var_to_counter_example var x _ r)))
Equations
- slim_check.unused_var_testable var α β = {run := λ (tracing min : bool), slim_check.testable.run β tracing min >>= λ (r : slim_check.test_result β), pure (slim_check.convert_counter_example _ r (psum.inr _))}
Equations
- slim_check.subtype_var_testable var var' α β p = {run := λ (tracing min : bool), slim_check.testable.run (∀ (x : subtype p), β x.val) tracing min >>= λ (r : slim_check.test_result (∀ (x : subtype p), β x.val)), pure (slim_check.convert_counter_example' _ r)}
Equations
- slim_check.decidable_testable p = {run := λ (tracing min : bool), return (dite p (λ (h : p), slim_check.test_result.success (psum.inr h)) (λ (h : ¬p), slim_check.test_result.failure h list.nil))}
Execute cmd
and repeat every time the result is gave_up
(at most
n
times).
Equations
- slim_check.retry cmd n.succ = cmd >>= λ (r : slim_check.test_result p), slim_check.retry._match_1 (slim_check.retry cmd n) r
- slim_check.retry cmd 0 = return (slim_check.test_result.gave_up 1)
- slim_check.retry._match_1 _f_1 (slim_check.test_result.failure Hce xs) = return (slim_check.test_result.failure Hce xs)
- slim_check.retry._match_1 _f_1 (slim_check.test_result.gave_up _x) = _f_1
- slim_check.retry._match_1 _f_1 (slim_check.test_result.success hp) = return (slim_check.test_result.success hp)
Count the number of times the test procedure gave up.
Equations
- slim_check.give_up x (slim_check.test_result.failure Hce xs) = slim_check.test_result.failure Hce xs
- slim_check.give_up x (slim_check.test_result.gave_up n) = slim_check.test_result.gave_up (n + x)
- slim_check.give_up x (slim_check.test_result.success (psum.inr p_1)) = slim_check.test_result.success (psum.inr p_1)
- slim_check.give_up x (slim_check.test_result.success (psum.inl ())) = slim_check.test_result.gave_up x
Try n
times to find a counter-example for p
.
Equations
- slim_check.testable.run_suite_aux p cfg r n.succ = let size : ℕ := (cfg.num_inst - n - 1) * cfg.max_size / cfg.num_inst in slim_check.retry ((slim_check.testable.run p cfg.enable_tracing tt).run {down := size}) 10 >>= λ (x : slim_check.test_result p), slim_check.testable.run_suite_aux._match_1 p (slim_check.testable.run_suite_aux p cfg r n) (λ (g : ℕ), slim_check.testable.run_suite_aux p cfg (slim_check.give_up g r) n) x
- slim_check.testable.run_suite_aux p cfg r 0 = return r
- slim_check.testable.run_suite_aux._match_1 p _f_1 _f_2 (slim_check.test_result.failure Hce xs) = return (slim_check.test_result.failure Hce xs)
- slim_check.testable.run_suite_aux._match_1 p _f_1 _f_2 (slim_check.test_result.gave_up g) = _f_2 g
- slim_check.testable.run_suite_aux._match_1 p _f_1 _f_2 (slim_check.test_result.success (psum.inr Hp)) = return (slim_check.test_result.success (psum.inr Hp))
- slim_check.testable.run_suite_aux._match_1 p _f_1 _f_2 (slim_check.test_result.success (psum.inl ())) = _f_1
Try to find a counter-example of p
.
Equations
Run a test suite for p
in io
.
Equations
- slim_check.testable.check' p cfg = slim_check.testable.check'._match_1 p cfg cfg.random_seed
- slim_check.testable.check'._match_1 p cfg (some seed) = io.run_rand_with seed (slim_check.testable.run_suite p cfg)
- slim_check.testable.check'._match_1 p cfg none = io.run_rand (slim_check.testable.run_suite p cfg)
Decorations
Instances of testable
use named_binder
as a decoration on
propositions in order to access the name of bound variables, as in
named_binder "x" (forall x, x < y)
. This helps the
testable
instances create useful error messages where variables
are matched with values that falsify a given proposition.
The following functions help support the gadget so that the user does not have to put them in themselves.
add_existential_decorations p
adds a
named_binderannotation at the
root of
pif
p` is an existential quantification.
Traverse the syntax of a proposition to find universal quantifiers
and existential quantifiers and add named_binder
annotations next to
them.
decorations_of p
is used as a hint to mk_decorations
to specify
that the goal should be satisfied with a proposition equivalent to p
with added annotations.
Equations
In a goal of the shape ⊢ tactic.decorations_of p
, mk_decoration
examines
the syntax of p
and add named_binder
around universal quantifications and
existential quantifications to improve error messages.
This tool can be used in the declaration of a function as follows:
def foo (p : Prop) (p' : tactic.decorations_of p . mk_decorations) [testable p'] : ...
p
is the parameter given by the user, p'
is an equivalent proposition where
the quantifiers are annotated with named_binder
.
Run a test suite for p
and return true or false: should we believe that p
holds?
Equations
- slim_check.testable.check p cfg p' = slim_check.testable.check._match_1 p cfg p' cfg.random_seed >>= λ (x : slim_check.test_result p'), slim_check.testable.check._match_2 p cfg p' x
- slim_check.testable.check._match_1 p cfg p' (some seed) = io.run_rand_with seed (slim_check.testable.run_suite p' cfg)
- slim_check.testable.check._match_1 p cfg p' none = io.run_rand (slim_check.testable.run_suite p' cfg)
- slim_check.testable.check._match_2 p cfg p' (slim_check.test_result.failure _x xs) = let counter_ex : string := "\n".intercalate xs in io.fail ("\n===================\nFound problems!\n\n" ++ to_string counter_ex ++ "\n-------------------\n")
- slim_check.testable.check._match_2 p cfg p' (slim_check.test_result.gave_up n) = io.fail ("Gave up " ++ to_string (repr n) ++ " times")
- slim_check.testable.check._match_2 p cfg p' (slim_check.test_result.success _x) = when (¬↥(cfg.quiet)) (io.put_str_ln "Success")