The M construction as a multivariate polynomial functor.
M types are potentially infinite tree-like structures. They are defined as the greatest fixpoint of a polynomial functor.
Main definitions
M.mk
- constructorM.dest
- destructorM.corec
- corecursor: useful for formulating infinite, productive computationsM.bisim
- bisimulation: proof technique to show the equality of infinite objects
Implementation notes
Dual view of M-types:
Mp
: polynomial functorM
: greatest fixed point of a polynomial functor
Specifically, we define the polynomial functor Mp
as:
- A := a possibly infinite tree-like structure without information in the nodes
- B := given the tree-like structure
t
,B t
is a valid path from the root oft
to any given node.
As a result Mp.obj α
is made of a dataless tree and a function from
its valid paths to values of α
The difference with the polynomial functor of an initial algebra is
that A
is a possibly infinite tree.
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)) (x : P.last.M) (a : P.A) (f : P.last.B a → P.last.M), x.dest = ⟨a, f⟩ → Π (i : fin2 n), P.drop.B a i → mvpfunctor.M.path P x i
- child : Π {n : ℕ} (P : mvpfunctor (n + 1)) (x : P.last.M) (a : P.A) (f : P.last.B a → P.last.M), x.dest = ⟨a, f⟩ → Π (j : P.last.B a) (i : fin2 n), mvpfunctor.M.path P (f j) i → mvpfunctor.M.path P x i
A path from the root of a tree to one of its node
Equations
- mvpfunctor.M.path.inhabited P x = {default := mvpfunctor.M.path.root x x.head x.children _ i (inhabited.default (P.drop.B x.head i))}
Polynomial functor of the M-type of P
. A
is a data-less
possibly infinite 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
n
-ary M-type for P
Equations
- P.mvfunctor_M = id P.Mp.mvfunctor
Equations
construct through corecursion the shape of an M-type without its contents
Equations
- mvpfunctor.M.corec_shape P g₀ g₂ = pfunctor.M.corec (λ (b : β), ⟨g₀ b, g₂ b⟩)
Proof of type equality as a function
Equations
- P.cast_lastB h = λ (b : P.last.B a), h.rec_on b
Using corecursion, construct the contents of an M-type
Equations
- mvpfunctor.M.corec_contents P g₀ g₁ g₂ x b h i (mvpfunctor.M.path.child x a f h' j i c) = have h₀ : a = g₀ b, from _, have h₁ : f j = mvpfunctor.M.corec_shape P g₀ g₂ (g₂ b (P.cast_lastB h₀ j)), from _, mvpfunctor.M.corec_contents P g₀ g₁ g₂ (f j) (g₂ b (P.cast_lastB h₀ j)) h₁ i c
- mvpfunctor.M.corec_contents P g₀ g₁ g₂ x b h i (mvpfunctor.M.path.root x a f h' i c) = have this : a = g₀ b, from _, g₁ b i (P.cast_dropB this i c)
Corecursor for M-type of P
Equations
- mvpfunctor.M.corec' P g₀ g₁ g₂ = λ (b : β), ⟨mvpfunctor.M.corec_shape P g₀ g₂ b, mvpfunctor.M.corec_contents P g₀ g₁ g₂ (mvpfunctor.M.corec_shape P g₀ g₂ b) b _⟩
Corecursor for M-type of P
Equations
- mvpfunctor.M.corec P g = mvpfunctor.M.corec' P (λ (b : β), (g b).fst) (λ (b : β), typevec.drop_fun (g b).snd) (λ (b : β), typevec.last_fun (g b).snd)
Implementation of destructor for M-type of P
Equations
- mvpfunctor.M.path_dest_left P h f' = λ (i : fin2 n) (c : P.drop.B a i), f' i (mvpfunctor.M.path.root x a f h i c)
Implementation of destructor for M-type of P
Equations
- mvpfunctor.M.path_dest_right P h f' = λ (j : P.last.B a) (i : fin2 n) (c : mvpfunctor.M.path P (f j) i), f' i (mvpfunctor.M.path.child x a f h j i c)
Destructor for M-type of P
Equations
- mvpfunctor.M.dest' P h f' = ⟨a, typevec.split_fun (mvpfunctor.M.path_dest_left P h f') (λ (x_1 : (P.B a).last), ⟨f x_1, mvpfunctor.M.path_dest_right P h f' x_1⟩)⟩
Destructor for M-types
Equations
- mvpfunctor.M.dest P x = mvpfunctor.M.dest' P _ x.snd
Constructor for M-types
Equations
- mvpfunctor.M.mk P = mvpfunctor.M.corec P (λ (i : P.obj (α ::: P.M α)), mvfunctor.map (typevec.id ::: mvpfunctor.M.dest P) i)