slim_check
: generators for functions #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file defines sampleable
instances for α → β
functions and
ℤ → ℤ
injective functions.
Functions are generated by creating a list of pairs and one more value using the list as a lookup table and resorting to the additional value when a value is not found in the table.
Injective functions are generated by creating a list of numbers and a permutation of that list. The permutation insures that every input is mapped to a unique output. When an input is not found in the list the input itself is used as an output.
Injective functions f : α → α
could be generated easily instead of
ℤ → ℤ
by generating a list α
, removing duplicates and creating a
permutations. One has to be careful when generating the domain to make
if vast enough that, when generating arguments to apply f
to,
they argument should be likely to lie in the domain of f
. This is
the reason that injective functions f : ℤ → ℤ
are generated by
fixing the domain to the range [-2*size .. -2*size]
, with size
the size parameter of the gen
monad.
Much of the machinery provided in this file is applicable to generate
injective functions of type α → α
and new instances should be easy
to define.
Other classes of functions such as monotone functions can generated using similar techniques. For monotone functions, generating two lists, sorting them and matching them should suffice, with appropriate default values. Some care must be taken for shrinking such functions to make sure their defining property is invariant through shrinking. Injective functions are an example of how complicated it can get.
- with_default : Π {α : Type u} {β : Type v}, list (Σ (_x : α), β) → β → slim_check.total_function α β
Data structure specifying a total function using a list of pairs and a default value returned when the input is not in the domain of the partial function.
with_default f y
encodes x ↦ f x
when x ∈ f
and x ↦ y
otherwise.
We use Σ
to encode mappings instead of ×
because we
rely on the association list API defined in data.list.sigma
.
Instances for slim_check.total_function
- slim_check.total_function.has_sizeof_inst
- slim_check.total_function.inhabited
- slim_check.total_function.has_repr
- slim_check.total_function.has_sizeof
Apply a total function to an argument.
Equations
- (slim_check.total_function.with_default m y).apply x = (list.lookup x m).get_or_else y
Implementation of has_repr (total_function α β)
.
Creates a string for a given finmap
and output, x₀ ↦ y₀, .. xₙ ↦ yₙ
for each of the entries. The brackets are provided by the calling function.
Equations
- slim_check.total_function.repr_aux m = string.join (list.qsort (λ (x y : string), decidable.to_bool (x < y)) (list.map (λ (x : Σ (_x : α), β), "" ++ to_string (repr x.fst) ++ (" ↦ " ++ to_string (repr x.snd) ++ ", ")) m))
Produce a string for a given total_function
.
The output is of the form [x₀ ↦ f x₀, .. xₙ ↦ f xₙ, _ ↦ y]
.
Equations
- (slim_check.total_function.with_default m y).repr = "[" ++ to_string (slim_check.total_function.repr_aux m) ++ ("_ ↦ " ++ to_string (has_repr.repr y) ++ "]")
Equations
- slim_check.total_function.has_repr α β = {repr := slim_check.total_function.repr _inst_2}
Redefine sizeof
to follow the structure of sampleable
instances.
Equations
Equations
Shrink a total function by shrinking the lists that represent it.
Equations
- slim_check.total_function.shrink (slim_check.total_function.with_default m x) = lazy_list.map (λ (_x : {y // sizeof y < sizeof (m, x)}), slim_check.total_function.shrink._match_1 m x _x) (slim_check.sampleable.shrink (m, x))
- slim_check.total_function.shrink._match_1 m x ⟨(m', x'), h⟩ = ⟨slim_check.total_function.with_default m'.dedupkeys x', _⟩
Equations
- slim_check.total_function.pi.sampleable_ext = {proxy_repr := slim_check.total_function α β, wf := slim_check.total_function.has_sizeof _inst_2, interp := slim_check.total_function.apply (λ (a b : α), _inst_3 a b), p_repr := slim_check.total_function.has_repr α β _inst_5, sample := slim_check.sampleable.sample (list (α × β)) >>= λ (xs : list (α × β)), uliftable.up (slim_check.sampleable.sample β) >>= λ (_p : ulift β), slim_check.total_function.pi.sampleable_ext._match_1 xs _p, shrink := slim_check.total_function.shrink (λ (a b : α), _inst_3 a b)}
- slim_check.total_function.pi.sampleable_ext._match_1 xs {down := x} = has_pure.pure (slim_check.total_function.with_default (slim_check.total_function.list.to_finmap' xs) x)
Map a total_function to one whose default value is zero so that it represents a finsupp.
Equations
The support of a zero default total_function
.
Equations
- (slim_check.total_function.with_default A y).zero_default_supp = (list.map sigma.fst (list.filter (λ (ab : Σ (_x : α), β), ab.snd ≠ 0) A.dedupkeys)).to_finset
Create a finitely supported function from a total function by taking the default value to zero.
Equations
- tf.apply_finsupp = {support := tf.zero_default_supp, to_fun := tf.zero_default.apply, mem_support_to_fun := _}
Equations
- slim_check.total_function.finsupp.sampleable_ext = {proxy_repr := slim_check.total_function α β, wf := slim_check.total_function.has_sizeof _inst_5, interp := slim_check.total_function.apply_finsupp (λ (a b : β), _inst_3 a b), p_repr := slim_check.total_function.has_repr α β _inst_7, sample := slim_check.sampleable.sample (list (α × β)) >>= λ (xs : list (α × β)), uliftable.up (slim_check.sampleable.sample β) >>= λ (_p : ulift β), slim_check.total_function.finsupp.sampleable_ext._match_1 xs _p, shrink := slim_check.total_function.shrink (λ (a b : α), _inst_2 a b)}
- slim_check.total_function.finsupp.sampleable_ext._match_1 xs {down := x} = has_pure.pure (slim_check.total_function.with_default (slim_check.total_function.list.to_finmap' xs) x)
Equations
- slim_check.total_function.dfinsupp.sampleable_ext = {proxy_repr := slim_check.total_function α β, wf := slim_check.total_function.has_sizeof _inst_5, interp := finsupp.to_dfinsupp ∘ slim_check.total_function.apply_finsupp, p_repr := slim_check.total_function.has_repr α β _inst_7, sample := slim_check.sampleable.sample (list (α × β)) >>= λ (xs : list (α × β)), uliftable.up (slim_check.sampleable.sample β) >>= λ (_p : ulift β), slim_check.total_function.dfinsupp.sampleable_ext._match_1 xs _p, shrink := slim_check.total_function.shrink (λ (a b : α), _inst_2 a b)}
- slim_check.total_function.dfinsupp.sampleable_ext._match_1 xs {down := x} = has_pure.pure (slim_check.total_function.with_default (slim_check.total_function.list.to_finmap' xs) x)
Equations
- slim_check.total_function.pi_pred.sampleable_ext = {proxy_repr := slim_check.sampleable_ext.proxy_repr (α → bool) _inst_1, wf := slim_check.sampleable_ext.wf _inst_1, interp := λ (m : slim_check.sampleable_ext.proxy_repr (α → bool)) (x : α), ↥(slim_check.sampleable_ext.interp (α → bool) m x), p_repr := slim_check.sampleable_ext.p_repr _inst_1, sample := slim_check.sampleable_ext.sample (α → bool) _inst_1, shrink := slim_check.sampleable_ext.shrink _inst_1}
Equations
- slim_check.total_function.pi_uncurry.sampleable_ext = {proxy_repr := slim_check.sampleable_ext.proxy_repr (α × β → γ) _inst_1, wf := slim_check.sampleable_ext.wf _inst_1, interp := λ (m : slim_check.sampleable_ext.proxy_repr (α × β → γ)) (x : α) (y : β), slim_check.sampleable_ext.interp (α × β → γ) m (x, y), p_repr := slim_check.sampleable_ext.p_repr _inst_1, sample := slim_check.sampleable_ext.sample (α × β → γ) _inst_1, shrink := slim_check.sampleable_ext.shrink _inst_1}
- map_to_self : Π {α : Type u} (xs : list (Σ (_x : α), α)), list.map sigma.fst xs ~ list.map sigma.snd xs → (list.map sigma.snd xs).nodup → slim_check.injective_function α
Data structure specifying a total function using a list of pairs and a default value returned when the input is not in the domain of the partial function.
map_to_self f
encodes x ↦ f x
when x ∈ f
and x ↦ x
,
i.e. x
to itself, otherwise.
We use Σ
to encode mappings instead of ×
because we
rely on the association list API defined in data.list.sigma
.
Instances for slim_check.injective_function
- slim_check.injective_function.has_sizeof_inst
- slim_check.injective_function.inhabited
- slim_check.injective_function.has_repr
- slim_check.injective_function.has_sizeof
Apply a total function to an argument.
Equations
- (slim_check.injective_function.map_to_self m _x _x_1).apply x = (list.lookup x m).get_or_else x
Produce a string for a given total_function
.
The output is of the form [x₀ ↦ f x₀, .. xₙ ↦ f xₙ, x ↦ x]
.
Unlike for total_function
, the default value is not a constant
but the identity function.
Equations
- (slim_check.injective_function.map_to_self m _x _x_1).repr = "[" ++ to_string (slim_check.total_function.repr_aux m) ++ "x ↦ x]"
Equations
Interpret a list of pairs as a total function, defaulting to the identity function when no entries are found for a given function
Equations
Remove a slice of length m
at index n
in a list and a permutation, maintaining the property
that it is a permutation.
A lazy list, in decreasing order, of sizes that should be
sliced off a list of length n
Equations
- slim_check.injective_function.slice_sizes n = dite (0 < n) (λ (h : 0 < n), lazy_list.cons ⟨n, h⟩ (λ («_» : unit), slim_check.injective_function.slice_sizes (n / 2))) (λ (h : ¬0 < n), lazy_list.nil)
Shrink a permutation of a list, slicing a segment in the middle.
The sizes of the slice being removed start at n
(with n
the length
of the list) and then n / 2
, then n / 4
, etc down to 1. The slices
will be taken at index 0
, n / k
, 2n / k
, 3n / k
, etc.
Equations
- slim_check.injective_function.shrink_perm xs = let k : ℕ := xs.fst.length in slim_check.injective_function.slice_sizes k >>= λ (n : ℕ+), lazy_list.of_list (list.fin_range (k / ↑n)) >>= λ (i : fin (k / ↑n)), have this : ↑i * ↑n < xs.fst.length, from _, has_pure.pure ⟨slim_check.injective_function.perm.slice (↑i * ↑n) ↑n xs, _⟩
Equations
- slim_check.injective_function.has_sizeof = {sizeof := λ (_x : slim_check.injective_function α), slim_check.injective_function.has_sizeof._match_1 _x}
- slim_check.injective_function.has_sizeof._match_1 (slim_check.injective_function.map_to_self xs _x _x_1) = sizeof (list.map sigma.fst xs)
Shrink an injective function slicing a segment in the middle of the domain and removing the corresponding elements in the codomain, hence maintaining the property that one is a permutation of the other.
Equations
- slim_check.injective_function.shrink (slim_check.injective_function.map_to_self xs h₀ h₁) = slim_check.injective_function.shrink_perm ⟨list.map sigma.fst xs, ⟨list.map sigma.snd xs, _⟩⟩ >>= λ (_p : {y // slim_check.sizeof_lt y ⟨list.map sigma.fst xs, ⟨list.map sigma.snd xs, _⟩⟩}), slim_check.injective_function.shrink._match_1 xs h₀ h₁ _p
- slim_check.injective_function.shrink._match_1 xs h₀ h₁ ⟨⟨xs', ⟨ys', _⟩⟩, h₂⟩ = have h₃ : xs'.length ≤ ys'.length, from _, have h₄ : ys'.length ≤ xs'.length, from _, has_pure.pure ⟨slim_check.injective_function.map_to_self (list.map prod.to_sigma (xs'.zip ys')) _ _, _⟩
Create an injective function from one list and a permutation of that list.
Equations
- slim_check.injective_function.mk xs ys h h' = have h₀ : xs.length ≤ ys.length, from _, have h₁ : ys.length ≤ xs.length, from _, slim_check.injective_function.map_to_self (slim_check.total_function.list.to_finmap' (xs.zip ys)) _ _
Equations
- slim_check.injective_function.pi_injective.sampleable_ext = {proxy_repr := slim_check.injective_function ℤ, wf := slim_check.injective_function.has_sizeof int.has_sizeof_inst, interp := λ (f : slim_check.injective_function ℤ), ⟨f.apply, _⟩, p_repr := slim_check.injective_function.has_repr ℤ int.has_repr, sample := slim_check.gen.sized (λ (sz : ℕ), let xs' : list ℤ := (-(2 * ↑sz + 2)).range (2 * ↑sz + 2) in slim_check.gen.permutation_of xs' >>= λ (ys : subtype xs'.perm), have Hinj : function.injective (λ (r : ℕ), -(2 * ↑sz + 2) + ↑r), from _, let r : slim_check.injective_function ℤ := slim_check.injective_function.mk xs' ys.val _ _ in has_pure.pure r), shrink := slim_check.injective_function.shrink (λ (a b : ℤ), int.decidable_eq a b)}