The W construction as a multivariate polynomial functor.
W types are well-founded tree-like structures. They are defined as the least fixpoint of a polynomial functor.
Main definitions
W_mk
- constructor- `W_dest - destructor
W_rec
- recursor: basis for defining functions by structural recursion onP.W α
W_rec_eq
- defining equation forW_rec
W_ind
- induction principle forP.W α
Implementation notes
Three views of M-types:
Wp
: polynomial functorW
: data type inductively defined by a triple: shape of the root, data in the root and children of the rootW
: least fixed point of a polynomial functor
Specifically, we define the polynomial functor Wp
as:
- A := a tree-like structure without information in the nodes
- B := given the tree-like structure
t
,B t
is a valid path (specified inductively byW_path
) from the root oft
to any given node.
As a result Wp.obj α
is made of a dataless tree and a function from
its valid paths to values of α
Reference
- [Jeremy Avigad, Mario M. Carneiro and Simon Hudon, Data Types as Quotients of Polynomial Functors][avigad-carneiro-hudon2019]
- root : Π {n : ℕ} (P : mvpfunctor (n + 1)) (a : P.A) (f : P.last.B a → P.last.W) (i : fin2 n), P.drop.B a i → P.W_path (W.mk a f) i
- child : Π {n : ℕ} (P : mvpfunctor (n + 1)) (a : P.A) (f : P.last.B a → P.last.W) (i : fin2 n) (j : P.last.B a), P.W_path (f j) i → P.W_path (W.mk a f) i
A path from the root of a tree to one of its node
Equations
- mvpfunctor.W_path.inhabited P x = {default := mvpfunctor.W_path.inhabited._match_1 P x I}
- mvpfunctor.W_path.inhabited._match_1 P (W.mk a f) I = mvpfunctor.W_path.root a f i (inhabited.default (P.drop.B (pfunctor.W.head (W.mk a f)) i))
Specialized destructor on W_path
Equations
- P.W_path_cases_on g' g = id (λ (i : fin2 n) (x : P.W_path (W.mk a f) i), x.cases_on (λ (x_a : P.A) (x_f : P.last.B x_a → P.last.W) (x_i : fin2 n) (x_c : P.drop.B x_a x_i) (H_1 : W.mk a f = W.mk x_a x_f), W.no_confusion H_1 (λ (a_eq : a = x_a), eq.rec (λ (x_f : P.last.B a → P.last.W) (x_c : P.drop.B a x_i) (f_eq : f == x_f), eq.rec (λ (H_2 : i = x_i), eq.rec (λ (x_c : P.drop.B a i) (H_3 : x == mvpfunctor.W_path.root a f i x_c), eq.rec ((λ (f : P.last.B a → P.last.W) (g' : (P.drop.B a).arrow α) (g : Π (j : P.last.B a), typevec.arrow (P.W_path (f j)) α) (i : fin2 n) (c : P.drop.B a i), g' i c) f g' g i x_c) _) H_2 x_c) _) a_eq x_f x_c)) (λ (x_a : P.A) (x_f : P.last.B x_a → P.last.W) (x_i : fin2 n) (x_j : P.last.B x_a) (x_c : P.W_path (x_f x_j) x_i) (H_1 : W.mk a f = W.mk x_a x_f), W.no_confusion H_1 (λ (a_eq : a = x_a), eq.rec (λ (x_f : P.last.B a → P.last.W) (x_j : P.last.B a) (x_c : P.W_path (x_f x_j) x_i) (f_eq : f == x_f), eq.rec (λ (x_c : P.W_path (f x_j) x_i) (H_2 : i = x_i), eq.rec (λ (x_c : P.W_path (f x_j) i) (H_3 : x == mvpfunctor.W_path.child a f i x_j x_c), eq.rec ((λ (f : P.last.B a → P.last.W) (g' : (P.drop.B a).arrow α) (g : Π (j : P.last.B a), typevec.arrow (P.W_path (f j)) α) (i : fin2 n) (j : P.last.B a) (c : P.W_path (f j) i), g j i c) f g' g i x_j x_c) _) H_2 x_c) _ x_c) a_eq x_f x_j x_c)) _ _ _)
Specialized destructor on W_path
Equations
- P.W_path_dest_left h = λ (i : fin2 n) (c : P.drop.B a i), h i (mvpfunctor.W_path.root a f i c)
Specialized destructor on W_path
Equations
- P.W_path_dest_right h = λ (j : P.last.B a) (i : fin2 n) (c : P.W_path (f j) i), h i (mvpfunctor.W_path.child a f i j c)
Polynomial functor for the W-type of P
. A
is a data-less well-founded
tree whereas, for a given a : A
, B a
is a valid path in tree a
so
that Wp.obj α
is made of a tree and a function from its valid paths to
the values it contains
W-type of P
Equations
- P.mvfunctor_W = id P.Wp.mvfunctor
First, describe operations on W
as a polynomial functor.
Recursor for Wp
Now think of W as defined inductively by the data ⟨a, f', f⟩ where
a : P.A
is the shape of the top nodef' : P.drop.B a ⟹ α
is the contents of the top nodef : P.last.B a → P.last.W
are the subtrees
Constructor for W
Recursor for W
Equations
- P.W_rec g ⟨a, f'⟩ = let g' : Π (a : P.A) (f : P.last.B a → P.last.W), typevec.arrow (P.W_path (W.mk a f)) α → (P.last.B a → C) → C := λ (a : P.A) (f : P.last.B a → P.last.W) (h : typevec.arrow (P.W_path (W.mk a f)) α) (h' : P.last.B a → C), g a (P.W_path_dest_left h) (λ (i : P.last.B a), ⟨f i, P.W_path_dest_right h i⟩) h' in P.Wp_rec g' a f'
W-types are functorial
Equations
- P.W_map g = λ (x : P.W α), mvfunctor.map g x
Constructor of a value of P.obj (α ::: β)
from components.
Useful to avoid complicated type annotation
Equations
- P.obj_append1 a f' f = ⟨a, typevec.split_fun f' f⟩
Yet another view of the W type: as a fixed point for a multivariate polynomial functor.
These are needed to use the W-construction to construct a fixed point of a qpf, since
the qpf axioms are expressed in terms of map
on P
.
Constructor for the W-type of P
Equations
- P.W_mk' ⟨a, f⟩ = P.W_mk a (typevec.drop_fun f) (typevec.last_fun f)
Destructor for the W-type of P