# W-types ```agda module trees.w-types where ``` <details><summary>Imports</summary> ```agda open import foundation.action-on-identifications-functions open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.empty-types open import foundation.equivalences open import foundation.function-extensionality open import foundation.function-types open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopies open import foundation.homotopy-induction open import foundation.identity-types open import foundation.propositional-truncations open import foundation.sets open import foundation.torsorial-type-families open import foundation.transport-along-identifications open import foundation.truncated-types open import foundation.truncation-levels open import foundation.type-theoretic-principle-of-choice open import foundation.universe-levels open import foundation.whiskering-homotopies open import trees.algebras-polynomial-endofunctors open import trees.coalgebras-polynomial-endofunctors open import trees.morphisms-algebras-polynomial-endofunctors open import trees.polynomial-endofunctors ``` </details> ## Idea Consider a type `A` equipped with a type family `B` over `A`. The type `W` generated inductively by a constructor `B x โ W` for each `x : A` is called the **W-type** `W A B` of `B`. The elements of `A` can be thought of as symbols for the constructors of `W A B`, and the functions `B x โ W A B` are the constructors. The elements of `W A B` can be thought of as well-founded trees. ## Definition ```agda data ๐ {l1 l2 : Level} (A : UU l1) (B : A โ UU l2) : UU (l1 โ l2) where tree-๐ : (x : A) (ฮฑ : B x โ ๐ A B) โ ๐ A B module _ {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} where shape-๐ : ๐ A B โ A shape-๐ (tree-๐ x ฮฑ) = x component-๐ : (x : ๐ A B) โ B (shape-๐ x) โ ๐ A B component-๐ (tree-๐ x ฮฑ) = ฮฑ ฮท-๐ : (x : ๐ A B) โ tree-๐ (shape-๐ x) (component-๐ x) ๏ผ x ฮท-๐ (tree-๐ x ฮฑ) = refl ``` ### W-types as algebras for a polynomial endofunctor ```agda structure-๐-Alg : {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} โ type-polynomial-endofunctor A B (๐ A B) โ ๐ A B structure-๐-Alg (pair x ฮฑ) = tree-๐ x ฮฑ ๐-Alg : {l1 l2 : Level} (A : UU l1) (B : A โ UU l2) โ algebra-polynomial-endofunctor (l1 โ l2) A B ๐-Alg A B = pair (๐ A B) structure-๐-Alg ``` ### W-types as coalgebras for a polynomial endofunctor ```agda ๐-Coalg : {l1 l2 : Level} (A : UU l1) (B : A โ UU l2) โ coalgebra-polynomial-endofunctor (l1 โ l2) A B pr1 (๐-Coalg A B) = ๐ A B pr1 (pr2 (๐-Coalg A B) x) = shape-๐ x pr2 (pr2 (๐-Coalg A B) x) = component-๐ x ``` ## Properties ### The elements of the form `tree-๐ x ฮฑ` where `B x` is an empty type are called the constants of `W A B` ```agda module _ {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} where constant-๐ : (x : A) โ is-empty (B x) โ ๐ A B constant-๐ x h = tree-๐ x (ex-falso โ h) is-constant-๐ : ๐ A B โ UU l2 is-constant-๐ x = is-empty (B (shape-๐ x)) ``` ### If each `B x` is inhabited, then the type `W A B` is empty ```agda module _ {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} where is-empty-๐ : ((x : A) โ type-trunc-Prop (B x)) โ is-empty (๐ A B) is-empty-๐ H (tree-๐ x ฮฑ) = apply-universal-property-trunc-Prop ( H x) ( empty-Prop) ( ฮป y โ is-empty-๐ H (ฮฑ y)) ``` ### Equality of W-types ```agda module _ {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} where Eq-๐ : ๐ A B โ ๐ A B โ UU (l1 โ l2) Eq-๐ (tree-๐ x ฮฑ) (tree-๐ y ฮฒ) = ฮฃ (x ๏ผ y) (ฮป p โ (z : B x) โ Eq-๐ (ฮฑ z) (ฮฒ (tr B p z))) refl-Eq-๐ : (w : ๐ A B) โ Eq-๐ w w refl-Eq-๐ (tree-๐ x ฮฑ) = pair refl (ฮป z โ refl-Eq-๐ (ฮฑ z)) center-total-Eq-๐ : (w : ๐ A B) โ ฮฃ (๐ A B) (Eq-๐ w) center-total-Eq-๐ w = pair w (refl-Eq-๐ w) aux-total-Eq-๐ : (x : A) (ฮฑ : B x โ ๐ A B) โ ฮฃ (B x โ ๐ A B) (ฮป ฮฒ โ (y : B x) โ Eq-๐ (ฮฑ y) (ฮฒ y)) โ ฮฃ (๐ A B) (Eq-๐ (tree-๐ x ฮฑ)) aux-total-Eq-๐ x ฮฑ (pair ฮฒ e) = pair (tree-๐ x ฮฒ) (pair refl e) contraction-total-Eq-๐ : (w : ๐ A B) (t : ฮฃ (๐ A B) (Eq-๐ w)) โ center-total-Eq-๐ w ๏ผ t contraction-total-Eq-๐ ( tree-๐ x ฮฑ) (pair (tree-๐ .x ฮฒ) (pair refl e)) = ap ( ( aux-total-Eq-๐ x ฮฑ) โ ( map-distributive-ฮ -ฮฃ { A = B x} { B = ฮป y โ ๐ A B} { C = ฮป y โ Eq-๐ (ฮฑ y)})) { x = ฮป y โ pair (ฮฑ y) (refl-Eq-๐ (ฮฑ y))} { y = ฮป y โ pair (ฮฒ y) (e y)} ( eq-htpy (ฮป y โ contraction-total-Eq-๐ (ฮฑ y) (pair (ฮฒ y) (e y)))) is-torsorial-Eq-๐ : (w : ๐ A B) โ is-torsorial (Eq-๐ w) is-torsorial-Eq-๐ w = pair (center-total-Eq-๐ w) (contraction-total-Eq-๐ w) Eq-๐-eq : (v w : ๐ A B) โ v ๏ผ w โ Eq-๐ v w Eq-๐-eq v .v refl = refl-Eq-๐ v is-equiv-Eq-๐-eq : (v w : ๐ A B) โ is-equiv (Eq-๐-eq v w) is-equiv-Eq-๐-eq v = fundamental-theorem-id ( is-torsorial-Eq-๐ v) ( Eq-๐-eq v) eq-Eq-๐ : (v w : ๐ A B) โ Eq-๐ v w โ v ๏ผ w eq-Eq-๐ v w = map-inv-is-equiv (is-equiv-Eq-๐-eq v w) equiv-Eq-๐-eq : (v w : ๐ A B) โ (v ๏ผ w) โ Eq-๐ v w equiv-Eq-๐-eq v w = pair (Eq-๐-eq v w) (is-equiv-Eq-๐-eq v w) is-trunc-๐ : (k : ๐) โ is-trunc (succ-๐ k) A โ is-trunc (succ-๐ k) (๐ A B) is-trunc-๐ k is-trunc-A (tree-๐ x ฮฑ) (tree-๐ y ฮฒ) = is-trunc-is-equiv k ( Eq-๐ (tree-๐ x ฮฑ) (tree-๐ y ฮฒ)) ( Eq-๐-eq (tree-๐ x ฮฑ) (tree-๐ y ฮฒ)) ( is-equiv-Eq-๐-eq (tree-๐ x ฮฑ) (tree-๐ y ฮฒ)) ( is-trunc-ฮฃ ( is-trunc-A x y) ( ฮป p โ is-trunc-ฮ k ( ฮป z โ is-trunc-is-equiv' k ( ฮฑ z ๏ผ ฮฒ (tr B p z)) ( Eq-๐-eq (ฮฑ z) (ฮฒ (tr B p z))) ( is-equiv-Eq-๐-eq (ฮฑ z) (ฮฒ (tr B p z))) ( is-trunc-๐ k is-trunc-A (ฮฑ z) (ฮฒ (tr B p z)))))) is-set-๐ : is-set A โ is-set (๐ A B) is-set-๐ = is-trunc-๐ neg-one-๐ ``` ### The structure map of the algebra `๐ A B` is an equivalence ```agda map-inv-structure-๐-Alg : {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} โ ๐ A B โ type-polynomial-endofunctor A B (๐ A B) map-inv-structure-๐-Alg (tree-๐ x ฮฑ) = pair x ฮฑ is-section-map-inv-structure-๐-Alg : {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} โ (structure-๐-Alg {B = B} โ map-inv-structure-๐-Alg {B = B}) ~ id is-section-map-inv-structure-๐-Alg (tree-๐ x ฮฑ) = refl is-retraction-map-inv-structure-๐-Alg : {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} โ (map-inv-structure-๐-Alg {B = B} โ structure-๐-Alg {B = B}) ~ id is-retraction-map-inv-structure-๐-Alg (pair x ฮฑ) = refl is-equiv-structure-๐-Alg : {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} โ is-equiv (structure-๐-Alg {B = B}) is-equiv-structure-๐-Alg = is-equiv-is-invertible map-inv-structure-๐-Alg is-section-map-inv-structure-๐-Alg is-retraction-map-inv-structure-๐-Alg equiv-structure-๐-Alg : {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} โ type-polynomial-endofunctor A B (๐ A B) โ ๐ A B equiv-structure-๐-Alg = pair structure-๐-Alg is-equiv-structure-๐-Alg is-equiv-map-inv-structure-๐-Alg : {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} โ is-equiv (map-inv-structure-๐-Alg {B = B}) is-equiv-map-inv-structure-๐-Alg = is-equiv-is-invertible structure-๐-Alg is-retraction-map-inv-structure-๐-Alg is-section-map-inv-structure-๐-Alg inv-equiv-structure-๐-Alg : {l1 l2 : Level} {A : UU l1} {B : A โ UU l2} โ ๐ A B โ type-polynomial-endofunctor A B (๐ A B) inv-equiv-structure-๐-Alg = pair map-inv-structure-๐-Alg is-equiv-map-inv-structure-๐-Alg ``` ### W-types are initial algebras for polynomial endofunctors ```agda map-hom-๐-Alg : {l1 l2 l3 : Level} {A : UU l1} {B : A โ UU l2} (X : algebra-polynomial-endofunctor l3 A B) โ ๐ A B โ type-algebra-polynomial-endofunctor X map-hom-๐-Alg X (tree-๐ x ฮฑ) = structure-algebra-polynomial-endofunctor X ( pair x (ฮป y โ map-hom-๐-Alg X (ฮฑ y))) structure-hom-๐-Alg : {l1 l2 l3 : Level} {A : UU l1} {B : A โ UU l2} (X : algebra-polynomial-endofunctor l3 A B) โ ( (map-hom-๐-Alg X) โ structure-๐-Alg) ~ ( ( structure-algebra-polynomial-endofunctor X) โ ( map-polynomial-endofunctor A B (map-hom-๐-Alg X))) structure-hom-๐-Alg X (pair x ฮฑ) = refl hom-๐-Alg : {l1 l2 l3 : Level} {A : UU l1} {B : A โ UU l2} (X : algebra-polynomial-endofunctor l3 A B) โ hom-algebra-polynomial-endofunctor (๐-Alg A B) X hom-๐-Alg X = pair (map-hom-๐-Alg X) (structure-hom-๐-Alg X) htpy-htpy-hom-๐-Alg : {l1 l2 l3 : Level} {A : UU l1} {B : A โ UU l2} (X : algebra-polynomial-endofunctor l3 A B) โ (f : hom-algebra-polynomial-endofunctor (๐-Alg A B) X) โ map-hom-๐-Alg X ~ map-hom-algebra-polynomial-endofunctor (๐-Alg A B) X f htpy-htpy-hom-๐-Alg {A = A} {B} X f (tree-๐ x ฮฑ) = ( ap ( ฮป t โ structure-algebra-polynomial-endofunctor X (pair x t)) ( eq-htpy (ฮป z โ htpy-htpy-hom-๐-Alg X f (ฮฑ z)))) โ ( inv ( structure-hom-algebra-polynomial-endofunctor (๐-Alg A B) X f ( pair x ฮฑ))) compute-structure-htpy-hom-๐-Alg : {l1 l2 l3 : Level} {A : UU l1} {B : A โ UU l2} (X : algebra-polynomial-endofunctor l3 A B) (x : A) (ฮฑ : B x โ ๐ A B) {f : ๐ A B โ type-algebra-polynomial-endofunctor X} โ (H : map-hom-๐-Alg X ~ f) โ ( ap ( structure-algebra-polynomial-endofunctor X) ( htpy-polynomial-endofunctor A B H (pair x ฮฑ))) ๏ผ ( ap ( ฮป t โ structure-algebra-polynomial-endofunctor X (pair x t)) ( eq-htpy (H ยทr ฮฑ))) compute-structure-htpy-hom-๐-Alg {A = A} {B} X x ฮฑ = ind-htpy ( map-hom-๐-Alg X) ( ฮป f H โ ( ap ( structure-algebra-polynomial-endofunctor X) ( htpy-polynomial-endofunctor A B H (pair x ฮฑ))) ๏ผ ( ap ( ฮป t โ structure-algebra-polynomial-endofunctor X (pair x t)) ( eq-htpy (H ยทr ฮฑ)))) ( ap ( ap (pr2 X)) ( coh-refl-htpy-polynomial-endofunctor A B ( map-hom-๐-Alg X) ( pair x ฮฑ)) โ ( inv ( ap ( ap (ฮป t โ pr2 X (pair x t))) ( eq-htpy-refl-htpy (map-hom-๐-Alg X โ ฮฑ))))) structure-htpy-hom-๐-Alg : {l1 l2 l3 : Level} {A : UU l1} {B : A โ UU l2} (X : algebra-polynomial-endofunctor l3 A B) โ (f : hom-algebra-polynomial-endofunctor (๐-Alg A B) X) โ ( structure-hom-๐-Alg X โh ( ( structure-algebra-polynomial-endofunctor X) ยทl ( htpy-polynomial-endofunctor A B (htpy-htpy-hom-๐-Alg X f)))) ~ ( ( (htpy-htpy-hom-๐-Alg X f) ยทr structure-๐-Alg {B = B}) โh ( structure-hom-algebra-polynomial-endofunctor (๐-Alg A B) X f)) structure-htpy-hom-๐-Alg {A = A} {B} X (pair f ฮผ-f) (pair x ฮฑ) = ( ( ( compute-structure-htpy-hom-๐-Alg X x ฮฑ ( htpy-htpy-hom-๐-Alg X (pair f ฮผ-f))) โ ( inv right-unit)) โ ( ap ( concat ( ap ( ฮป t โ pr2 X (pair x t)) ( eq-htpy (htpy-htpy-hom-๐-Alg X (pair f ฮผ-f) ยทr ฮฑ))) ( pr2 X (map-polynomial-endofunctor A B f (pair x ฮฑ)))) ( inv (left-inv ( ฮผ-f (pair x ฮฑ)))))) โ ( inv ( assoc ( ap ( ฮป t โ pr2 X (pair x t)) ( eq-htpy (htpy-htpy-hom-๐-Alg X (pair f ฮผ-f) ยทr ฮฑ))) ( inv (ฮผ-f (pair x ฮฑ))) ( ฮผ-f (pair x ฮฑ)))) htpy-hom-๐-Alg : {l1 l2 l3 : Level} {A : UU l1} {B : A โ UU l2} (X : algebra-polynomial-endofunctor l3 A B) โ (f : hom-algebra-polynomial-endofunctor (๐-Alg A B) X) โ htpy-hom-algebra-polynomial-endofunctor (๐-Alg A B) X (hom-๐-Alg X) f htpy-hom-๐-Alg X f = pair (htpy-htpy-hom-๐-Alg X f) (structure-htpy-hom-๐-Alg X f) is-initial-๐-Alg : {l1 l2 l3 : Level} {A : UU l1} {B : A โ UU l2} (X : algebra-polynomial-endofunctor l3 A B) โ is-contr (hom-algebra-polynomial-endofunctor (๐-Alg A B) X) is-initial-๐-Alg {A = A} {B} X = pair ( hom-๐-Alg X) ( ฮป f โ eq-htpy-hom-algebra-polynomial-endofunctor (๐-Alg A B) X (hom-๐-Alg X) f ( htpy-hom-๐-Alg X f)) ```