AmΓ©lia Liao
Open source, combination wiki (Γ la nLab) + Agda library.
open import 1Lab.Path
contains basic results about identity.
Very basic results in algebra + synthetic homotopy theory (e.g.Β HoTT ch.Β 8)
This talk is a module in the 1Lab!
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
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?
When are functions f,g:AβB identical?
Expected answer: whenever f(x)β‘g(x) for all x;
Actual answer: π€·ββοΈ!
Extensionality is independent of MLTT. E.g. antifunext.
But itβs for doing maths. Our solution: the interval, paths.
A type I
, with endpoints
i0
, i1
, which
(co)represents identifications.
An identification p:xβ‘Aβy is an element i:I,p(i):A that satisfies p(i0)=x and p(i1)=y.
In Agda: path lambdas and path application.
_ : {A : Type} (x : A) β x β‘ x _ = Ξ» x β Ξ» i β x
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
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.
Ideally: the identity type xβ‘Aβy computes to something simpler, based on the structure of A.
Observational type theory with strict propositionsβ¦
Higher observational type theoryβ¦
Key example: group homomorphisms
f,g:GβH.
These are pairs, so
fβ‘g
computes to a pair type.
But the second component β the proof that
f
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
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.
Same setup: when are types
A,B:Type
identical?
Ideal answer: when they are indistinguishable.
What distinguishes types?
Take e.g.Β NΓString vs.Β StringΓN:
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 (Aβ‘B) is equivalent to (AβB):
Just as much control over p:Aβ‘B as we would have over f:AβB;
Funny looking consequences: e.g.Β ZβN, and βZ is a groupβ, so transport lets us conclude βN is a group.β
Implications for library design?
The traditional approach for doing algebra in a proof assistant is by letting the elaborator fill in the structure
Technically different approaches, but, in effect, you can say x,y:Z and write βx+yβ to mean βthe addition on Zβ
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
:
Pod
is definitionally
P,
but type class search doesnβt respect that
Requires careful setup to avoid diamonds
Mostly sensible for concrete types, but βtheβ order on XΓY 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:
Type-theoretically, we can say:
These sometimes conflict: once we fix the multiplication, a monoid hasβ¦
Simplifying a bit, a monoid M is a tuple (M0β,β,1,Ξ»,Ο,Ξ±) consisting of
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)
is-monoid-is-prop _ _ _ = Ξ£-is-hlevel 1 (is-hlevel-is-prop 2) Ξ» mset β let instance _ = hlevel-instance {n = 2} mset in hlevel 1
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 M=(M0β,m,u,Ξ±) and N=(N0β,n,v,Ξ²), what is Mβ‘N?
module _ (M@(Mβ , m , u , Ξ±) N@(Nβ , n , v , Ξ²) : Monoid) where
Step 1: an identity of tuples is a tuple of identities, and identity
in is-monoid
is trivial;
Stepβ : Typeβ Stepβ = Ξ£[ p β Mβ β‘ Nβ ] ( PathP (Ξ» i β p i β p i β p i) m n Γ PathP (Ξ» i β p i) u v ) stepβ : Stepβ β M β‘ N stepβ (p , q , r) i = p i , q i , r i , is-propβpathp (Ξ» i β is-monoid-is-prop (p i) (q i) (r i)) Ξ± Ξ² i
Step 2: we can simplify the PathP
s 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!
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 D equipped with a βprojectionβ functor Ο:DβC; 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 EC, we can put together a category with objects βx:CβOb[x] β the total category β«E.
Similarly, each x:C gives a category Eβ(x) with objects Ob[x] and maps Hom[id](β,β) β the fibre over x.
Properties of the functor Ο:β«EβC map very nicely to properties of E!
Ο isβ¦ | E satisfiesβ¦ |
---|---|
faithful | each Hom[β](β,β) is a proposition |
full | each
Hom[β](β,β)
is inhabited |
fully faithful | each Hom[β](β,β) is contractible |
amnestic | each Eβ(x) univalent, and so is C |
That last row is important!
If we start with a functor, then Ob[x] is the type βy:DβΟ(y)β x; the βfibreβ of Ο over x.
private module _ {o β o' β' : Level} {C : Precategory o β} {D : Precategory o' β'} (Ο : Functor D C) where private private module C = Cat C module D = Cat D module Ο = Functor Ο open C using (_β_ ; id ; _β _) open C._β _ open Displayed
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.
These data amount to:
A type family of structures F:TypeβType;
A proposition Hom[f](S,T), given f:AβB, S:F(A), T:F(B).
This type defines what it means for f to be a βF-structure-preserving map from S to Tβ
Proofs that (2) contains the identities and is closed under composition;
For univalence: a proof that if id is a structure preserving map SβT (and also TβS), then S=T.
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
We can derive that this is a property pretty mechanically:
module _ where private unquoteDecl eqv = declare-record-iso eqv (quote is-semigroup) is-semigroup-is-prop : {A : Type β} {s : A β A β A} β is-prop (is-semigroup s) is-semigroup-is-prop = Isoβis-hlevel 1 eqv $ Ξ£-is-hlevel 1 (hlevel 1) Ξ» x β let instance _ = hlevel-instance {n = 2} x in hlevel 1
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
private unquoteDecl eqv = declare-record-iso eqv (quote Semigroup-on) unquoteDecl eqv' = declare-record-iso eqv' (quote is-semigroup-hom) instance H-Level-is-semigroup : β {n} {s : A β A β A} β H-Level (is-semigroup s) (suc n) H-Level-is-semigroup = prop-instance is-semigroup-is-prop H-Level-is-semigroup-hom : β {n} {f : A β B} {S : Semigroup-on A} {T : Semigroup-on B} β H-Level (is-semigroup-hom f S T) (suc n) H-Level-is-semigroup-hom {T = T} = prop-instance (Isoβis-hlevel 1 eqv' (hlevel 1)) where instance _ = hlevel-instance {n = 2} (T .Semigroup-on.has-is-set) open is-semigroup-hom
We can then fill in the four bullet points.
F
is Semigroup-on
,
Hom[β](β,β)
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
Cubical type theory works in practice, but handling the raw primitives bites
Univalence challenges us to reconsider the notion of structure