Cubical types for the working formaliserπŸ”—

AmΓ©lia Liao

The 1LabπŸ”—

Open source, combination wiki (Γ  la nLab) + Agda library.

  • Reimplementations of everything we need, e.g.
  open import 1Lab.Path

contains basic results about identity.

This talk is a module in the 1Lab!

The 1Lab as a libraryπŸ”—

  • Constructive but impredicative (propositional resizing); homotopical features used freely

  • Type classes for automation, but only of properties; equipping a type with structure (e.g.Β algebraic) is always explicit.

  • Playground for new Agda features, e.g.Β opaque, OVERLAPPING instances

This talkπŸ”—

  • HoTT simplifies Martin-LΓΆf type theory by giving universes a nice universal property (univalence)

  • Cubical type theory complicates type theory to make this compute (uaΞ²)

Lots of work has gone into usability of traditional proof assistants; but what about higher dimensional proof assistants?

Extensional equalityπŸ”—

When are functions identical?

  • Expected answer: whenever for all

  • Actual answer: πŸ€·β€β™€οΈ!

    Extensionality is independent of MLTT. E.g. antifunext.

But it’s for doing maths. Our solution: the interval, paths.

_ = I
_ = i0
_ = i1

A type I, with endpoints i0, i1, which (co)represents identifications.

An identification is an element that satisfies and

In Agda: path lambdas and path application.

_ : {A : Type} (x : A) β†’ x ≑ x
_ = Ξ» x β†’ Ξ» i β†’ x

Working with the intervalπŸ”—

  cong : (f : A β†’ B) (x y : A) β†’ x ≑ y β†’ f x ≑ f y
  cong f x y p = Ξ» i β†’ f (p i)


_ : {f : A β†’ B} β†’ funext {f = f} (Ξ» x β†’ refl) ≑ refl
_ = refl

_ : {f g : A β†’ B} {h : βˆ€ x β†’ f x ≑ g x}
  β†’ funext (Ξ» x β†’ sym (h x)) ≑ sym (funext h)
_ = refl

_ : {f g h : A β†’ B} {Ξ± : βˆ€ x β†’ f x ≑ g x} {Ξ² : βˆ€ x β†’ g x ≑ h x}
  β†’ funext (Ξ» x β†’ Ξ± x βˆ™ Ξ² x) ≑ funext Ξ± βˆ™ funext Ξ²
_ = refl

We can even explain things like pattern matching on HITs:

  example : (p : SΒΉ) β†’ SΒΉ
  example base = base
  example (loop i) = {! !}
  -- Goal: SΒΉ
  --  i = i0 ⊒ base
  --  i = i1 ⊒ base


It’s not all perfectπŸ”—

Introducing path abstractions does a number on inference:

  what : ((x y : A) β†’ x ≑ y) β†’ (x y : A) β†’ x ≑ y
  what h x y i = h _ _ i

Agda generates constraints:

  _x h x y (i = i0) = x
  _y h x y (i = i1) = y

Neither of these is fully determined!
Both metas go unsolved.

Need a nice interface for extensional equality.

A nice interfaceπŸ”—

Ideally: the identity type computes to something simpler, based on the structure of

  • Observational type theory with strict propositions…

    • Pro: computed identity type is basically always right;
    • Con: doesn’t scale to homotopy.
  • Higher observational type theory…

    • Pro: scales to homotopy;
    • Con: computed identity system will in many cases be overcomplicated.

Key example: group homomorphisms These are pairs, so computes to a pair type.
But the second component β€” the proof that is a homomorphism β€” is irrelevant!

Without a sort of strict propositions, the system can’t β€œsee” this.
So β€œprimitive higher-OTT” will still need a helper, to fill in the trivial second component.

Surprisingly, we can use type classes!

  record Extensional (A : Type) : Type (lsuc lzero) where
    field
      _β‰ˆ_     : A β†’ A β†’ Type
      to-path : βˆ€ {x y : A} β†’ x β‰ˆ y β†’ x ≑ y

And an overlappable instance for the base case:

    ext-base : Extensional A
    ext-base = record { to-path = Ξ» x β†’ x }
    {-# OVERLAPPABLE ext-base #-}

We can test that this works by asking Agda to check ext at various types:

  _ : {A B : Type} {f g : A β†’ B} β†’ (βˆ€ x β†’ f x ≑ g x) β†’ f ≑ g
  _ = ext

  _ : {A B C : Type} {f g : A β†’ B β†’ C} β†’ (βˆ€ x y β†’ f x y ≑ g x y) β†’ f ≑ g
  _ = ext

  _ : {A : Type} {x y : A} β†’ x ≑ y β†’ x ≑ y
  _ = ext

A benefit: type class solving isn’t a real function. Can be unstable under substitution!

  instance
    ext-uncurry : ⦃ Extensional (A β†’ B β†’ C) ⦄ β†’ Extensional (A Γ— B β†’ C)
    ext-uncurry ⦃ e ⦄ = record
      { _β‰ˆ_     = Ξ» f g β†’ e ._β‰ˆ_ (curry f) (curry g)
      ; to-path = Ξ» h i (x , y) β†’ e .to-path h i x y
      }
    {-# OVERLAPPING ext-uncurry #-}

The resulting relation has three arguments, rather than two:

  _ : {A B C D : Type} {f g : A Γ— B β†’ C β†’ D}
    β†’ (βˆ€ a b c β†’ f (a , b) c ≑ g (a , b) c)
    β†’ f ≑ g
  _ = ext

The real implementation handles maps of algebraic structures, e.g. groups,

  _ : {G H : Group lzero} {f g : Groups.Hom G H}
    β†’ ((x : ⌞ G ⌟) β†’ f # x ≑ g # x) β†’ f ≑ g
  _ = ext

maps from certain higher inductive types, e.g.Β generic set-quotients, or abelian group homomorphisms from a tensor product,

  _ : {A B C : Abelian-group lzero} {f g : Ab.Hom (A βŠ— B) C}
    β†’ ((x : ⌞ A ⌟) (y : ⌞ B ⌟) β†’ f # (x , y) ≑ g # (x , y))
    β†’ f ≑ g
  _ = ext

and also: natural transformations, maps in slice categories, in comma categories, in categories of elements, in wide and full subcategories, in categories of monad algebras, type-theoretic equivalences and embeddings, monotone maps, &c., &c.

Structure (and) identityπŸ”—

Same setup: when are types identical?
Ideal answer: when they are indistinguishable.
What distinguishes types?

Take e.g.Β  vs.Β 

  • In base MLTT (or even e.g.Β Lean): no proof that they’re distinct

  • … because if you give me any property that holds of one, I can modify your proof to show it holds of the other (by hand)

  • … so these types should be identical!

The actual answer: πŸ€·β€β™€οΈ
Univalence (and equality reflection!) says they’re the same; (set-level) OTT says they’re different; everyone else is undecided

Univalence says is equivalent to

  • Just as much control over as we would have over

    • E.g. exactly two inhabitants in

Funny looking consequences: e.g.Β  and β€œ is a group”, so transport lets us conclude β€œ is a group.”

Implications for library design?

Stuff vs.Β structure vs.Β propertyπŸ”—

The traditional approach for doing algebra in a proof assistant is by letting the elaborator fill in the structure

  • E.g. mathlib4 (Lean): extensive use of type classes
  • E.g. mathcomp (Rocq): canonical structures

Technically different approaches, but, in effect, you can say and write β€œβ€ to mean β€œthe addition on ”

From a user’s POV, great! Can write maths β€œas on paper” and the system does β€œthe right thing”.

But it’s actually pretty funny if you think about it:

  • β€œAdditive” monoid vs β€œmultiplicative” monoid; the same algebraic structure, but the entire hierarchy is duplicated

  • Type synonyms like OrderDual: is definitionally but type class search doesn’t respect that

  • Requires careful setup to avoid diamonds

  • Mostly sensible for concrete types, but β€œthe” order on is a choice!

In effect, mathematical vernacular makes explicit the stuff, leaves the structure implicit, and hardly ever mentions the properties.

Our approach is to always bundle types with the structure under consideration. Three-layered approach:

record is-ring {A : Type} (1r : A) (_*_ _+_ : A β†’ A β†’ A) : Type where
  -- _*_ is a monoid, _+_ is an abelian group, distributivity

record Ring-on (A : Type) : Type where
  field
    1r      : A
    _*_ _+_ : A β†’ A β†’ A

    has-is-ring : is-ring 1r _*_ _+_

Ring : Type₁
Ring = Σ[ T ∈ Type ] Ring-on T -- almost

Similarly is-group/Group-on/Group, is-monoid/Monoid-on/Monoid, etc.

This layered approach still requires a bit of choosing. Categorically, the guidelines are:

  • Forgetting the properties should be fully faithful
  • Forgetting the structure should be faithful

Type-theoretically, we can say:

  • The properties should be a family of propositions
  • The structure should preserve h-level

These sometimes conflict: once we fix the multiplication, a monoid has…

  • exactly one identity element (so we could put it in the properties), but
  • multiplication-preserving maps of monoids don’t necessarily preserve the identity (so it’s actually a structure).

Identity of structuresπŸ”—

Simplifying a bit, a monoid is a tuple consisting of

  • The stuff:
  • The structure: the binary operator and the identity;
  • The property: the proofs of identity on the left, on the right, and of associativity.

When are two monoids the same?

Let’s make it a bit more formal: we define the property, parametrised over the stuff and the structure;

is-monoid : (M : Type) β†’ (M β†’ M β†’ M) β†’ M β†’ Type
is-monoid M m u =
  is-set M            -- required so is-monoid is a property
  Γ— (βˆ€ x β†’ m u x ≑ x) -- left identity
  Γ— (βˆ€ x β†’ m x u ≑ x) -- right identity
  Γ— (βˆ€ x y z β†’ m x (m y z) ≑ m (m x y) z) -- associativity

We’ll need that it’s an actual property:

is-monoid-is-prop : βˆ€ M m u β†’ is-prop (is-monoid M m u)

Then we sum it all:

Monoid : Type₁
Monoid =
  Σ[ M ∈ Type ]                   -- stuff
  Ξ£[ m ∈ (M β†’ M β†’ M) ] Ξ£[ u ∈ M ] -- structure
    (is-monoid M m u)             -- property

Completely formally, a Monoid has 7 fields, but we’ve shown that 4 don’t matter.

Then we can compute the identity type: If we have monoids and what is

module _ (M@(Mβ‚€ , m , u , Ξ±) N@(Nβ‚€ , n , v , Ξ²) : Monoid) where

Step 2: we can simplify the PathPs to talk about transport instead:

  Stepβ‚‚ : Type₁
  Stepβ‚‚ = Ξ£[ p ∈ Mβ‚€ ≑ Nβ‚€ ]
    ( (βˆ€ x y β†’ transport p (m x y) ≑ n (transport p x) (transport p y))
    Γ— transport p u ≑ v
    )

  stepβ‚‚ : Stepβ‚‚ β†’ Step₁
  stepβ‚‚ (p , q , r) = p , q' , to-pathp r where
    -- q' actually pretty complicated...
    q' = funext-dep Ξ» Ξ± β†’ funext-dep Ξ» Ξ² β†’
      to-pathp (q _ _ βˆ™ apβ‚‚ n (from-pathp Ξ±) (from-pathp Ξ²))

Step 3: we know what identity of types is!

  Step₃ : Type
  Step₃ = Ξ£[ p ∈ Mβ‚€ ≃ Nβ‚€ ] -- equivalence!
    ( (βˆ€ x y β†’ p .fst (m x y) ≑ n (p .fst x) (p .fst y))
    Γ— p .fst u ≑ v
    )

Takes a bit of massaging…

  step₃ : Step₃ β†’ Stepβ‚‚
  step₃ (p , q , r) = ua p , q' , r' where
    r' = transport (ua p) u β‰‘βŸ¨ uaΞ² p u βŸ©β‰‘
         p .fst u           β‰‘βŸ¨ r βŸ©β‰‘
         v                  ∎

    q' : βˆ€ x y β†’ _
    q' x y =
      transport (ua p) (m x y)                    β‰‘βŸ¨ uaΞ² p (m x y) βŸ©β‰‘
      p .fst (m x y)                              β‰‘βŸ¨ q x y βŸ©β‰‘
      n (p .fst x) (p .fst y)                     β‰‘βŸ¨ apβ‚‚ n (sym (uaΞ² p x)) (sym (uaΞ² p y)) βŸ©β‰‘
      n (transport (ua p) x) (transport (ua p) y) ∎

The conclusion: identity in the type Monoid is exactlyΒΉ monoid isomorphism!

Designing for structure identityπŸ”—

While it’s great that the theory works, the proofs are pretty annoying.
But they’re very mechanical β€” incremental!

Our solution: we can make the three-layer approach from before into an actual theorem, using displayed categories.

An alternative presentation of the data of a category equipped with a β€œprojection” functor just more β€œnative” to type theory.

The idea: turn fibers into families.

  record Displayed o' β„“' : Type (o βŠ” β„“ βŠ” lsuc (o' βŠ” β„“')) where
    field
      Ob[_]  : ⌞ C ⌟ β†’ Type o'
      Hom[_] : (f : Hom x y) β†’ Ob[ x ] β†’ Ob[ y ] β†’ Type β„“'

Stuff over stuff, structure over structure:

      id'  : {x' : Ob[ x ]} β†’ Hom[ id ] x' x'
      _∘'_
        : {x' : Ob[ x ]} {y' : Ob[ y ]} {z' : Ob[ z ]}
        β†’ {f : Hom y z} (f' : Hom[ f ] y' z')
        β†’ {g : Hom x y} (g' : Hom[ g ] x' y')
        β†’ Hom[ f ∘ g ] x' z'

… also need property over property; check the formalisation.

If we start with a displayed category we can put together a category with objects β€” the total category

Similarly, each gives a category with objects and maps β€” the fibre over

Properties of the functor map very nicely to properties of

is… satisfies…
faithful each is a proposition
full each is inhabited
fully faithful each is contractible
amnestic each univalent, and so is

That last row is important!

If we start with a functor, then is the type the β€œfibre” of over

  Street : Displayed C _ _
  Street .Ob[_] x = Ξ£[ y ∈ D ] (Ο€.β‚€ y β‰… x)

The maps over are more complicated:

  Street .Hom[_] f (x' , i) (y' , j) = Σ[ g ∈ D.Hom x' y' ]
    (Ο€.₁ g ≑ j .from ∘ f ∘ i .to)

Accordingly, the structure is pretty annoying:

  Street .id' {x} {x' , i} = record { snd =
    Ο€.₁ D.id               β‰‘βŸ¨ Ο€.F-id βŸ©β‰‘
    id                     β‰‘Λ˜βŸ¨ i .invr βŸ©β‰‘Λ˜
    i .from ∘ i .to        β‰‘Λ˜βŸ¨ apβ‚‚ _∘_ refl (C.idl _) βŸ©β‰‘Λ˜
    i .from ∘ id C.∘ i .to ∎ }
  Street ._∘'_ = _ -- even worse!

We can present a concrete, univalent category as a displayed category.

  • Concrete: the base is and each is a proposition.
  • Univalent: each is a partial order.

These data amount to:

  1. A type family of structures

  2. A proposition given

    This type defines what it means for to be a β€œ map from to ”

  3. Proofs that (2) contains the identities and is closed under composition;

  4. For univalence: a proof that if is a structure preserving map (and also then

Concrete exampleπŸ”—

We start by defining the property, parametrised by the structure:

record is-semigroup {A : Type β„“} (_⋆_ : A β†’ A β†’ A) : Type β„“ where
  field
    has-is-set : is-set A
    assoc      : βˆ€ {x y z} β†’ x ⋆ (y ⋆ z) ≑ (x ⋆ y) ⋆ z

Then we define the structure:

record Semigroup-on (A : Type β„“) : Type β„“ where
  field
    _⋆_              : A β†’ A β†’ A
    has-is-semigroup : is-semigroup _⋆_
  open is-semigroup has-is-semigroup public

… and what it means for functions to preserve it:

record is-semigroup-hom (f : A β†’ B) (S : Semigroup-on A) (T : Semigroup-on B)
  : Type (level-of A βŠ” level-of B) where

  private
    module S = Semigroup-on S
    module T = Semigroup-on T

  field
    pres-⋆ : βˆ€ x y β†’ f (x S.⋆ y) ≑ f x T.⋆ f y

We can then fill in the four bullet points. is Semigroup-on, is is-semigroup-hom,

Semigroup-structure : βˆ€ {β„“} β†’ Thin-structure β„“ Semigroup-on
Semigroup-structure .is-hom f S T = el! (is-semigroup-hom f S T)

… identities and composites are respected …

Semigroup-structure .id-is-hom .pres-⋆ x y = refl
Semigroup-structure .∘-is-hom f g Ξ± Ξ² .pres-⋆ x y =
  ap f (Ξ² .pres-⋆ x y) βˆ™ Ξ± .pres-⋆ _ _

and, finally, the univalence condition:

Semigroup-structure .id-hom-unique p q = Iso.injective eqv $ Ξ£-prop-path! $
  ext Ξ» a b β†’ p .pres-⋆ a b

ConclusionπŸ”—

  • Cubical type theory works in practice, but handling the raw primitives bites

    • … but it is possible to do better
  • Univalence challenges us to reconsider the notion of structure

    • … but provides insights on how to organise mathematical libraries