Categorical models of homotopy type theory

Similar documents
Categorical Homotopy Type Theory

Algebraic models of homotopy type theory

Introduction to type theory and homotopy theory

1 / A bird s-eye view of type theory. 2 A bird s-eye view of homotopy theory. 3 Path spaces and identity types. 4 Homotopy type theory

Homotopy type theory: a new connection between logic, category theory and topology

WHAT IS AN ELEMENTARY HIGHER TOPOS?

The synthetic theory of -categories vs the synthetic theory of -categories

Abstracting away from cell complexes

Inductive and higher inductive types

Inductive and higher inductive types

Homotopy theory in type theory

Homotopy theory in type theory

in path component sheaves, and the diagrams

Internalizing the simplicial model structure in Homotopy Type Theory. Simon Boulier

Semantics and syntax of higher inductive types

Univalent Foundations and Set Theory

Algebraic model structures

T -spectra. Rick Jardine. March 2, University of Western Ontario

Higher toposes Internal logic Modalities Sub- -toposes Formalization. Modalities in HoTT. Egbert Rijke, Mike Shulman, Bas Spitters 1706.

Homotopy type theory: towards Grothendieck s dream

Cellularity, composition, and morphisms of algebraic weak factorization systems

Lecture 9: Sheaves. February 11, 2018

Local higher category theory

LOOP SPACES IN MOTIVIC HOMOTOPY THEORY. A Dissertation MARVIN GLEN DECKER

SEMANTICS OF HIGHER INDUCTIVE TYPES

The equivalence axiom and univalent models of type theory.

Dependent type theory

Homotopy Type Theory An Overview (unfinished)

Algebraic models for higher categories

ON THE HOMOTOPY THEORY OF ENRICHED CATEGORIES

Journal of Pure and Applied Algebra

Commutative ring objects in pro-categories and generalized Moore spectra

Recent progress in Homotopy type theory

Derived Algebraic Geometry III: Commutative Algebra

THE SIMPLICIAL MODEL OF UNIVALENT FOUNDATIONS

Graduate algebraic K-theory seminar

ON THE HOMOTOPY THEORY OF ENRICHED CATEGORIES

A Grothendieck site is a small category C equipped with a Grothendieck topology T. A Grothendieck topology T consists of a collection of subfunctors

Homotopy Theoretic Aspects of Constructive Type Theory

Homotopy Type Theory

University of Oxford, Michaelis November 16, Categorical Semantics and Topos Theory Homotopy type theor

ON THE COFIBRANT GENERATION OF MODEL CATEGORIES arxiv: v1 [math.at] 16 Jul 2009

Algebraic models for higher categories

Internal Language of Finitely Complete (, 1)-categories

Commutative ring objects in pro-categories and generalized Moore spectra

Derivatives of the identity functor and operads

Universität Regensburg Mathematik

Real-cohesion: from connectedness to continuity

A model-independent theory of -categories

An Algebraic Weak Factorisation System on 01-Substitution Sets: A Constructive Proof

ON LEFT AND RIGHT MODEL CATEGORIES AND LEFT AND RIGHT BOUSFIELD LOCALIZATIONS

Category Theory. Travis Dirle. December 12, 2017

POSTNIKOV EXTENSIONS OF RING SPECTRA

AXIOMS FOR GENERALIZED FARRELL-TATE COHOMOLOGY

REPRESENTATIONS OF SPACES

THE HOMOTOPY CALCULUS OF CATEGORIES AND GRAPHS

PART III.3. IND-COHERENT SHEAVES ON IND-INF-SCHEMES

Derived Algebraic Geometry I: Stable -Categories

Sheaf models of type theory

Postnikov extensions of ring spectra

Some glances at topos theory. Francis Borceux

MODEL STRUCTURES ON PRO-CATEGORIES

Outline of Synthetic Differential Geometry F. William Lawvere

Constructive Set Theory from a Weak Tarski Universe

28 The fundamental groupoid, revisited The Serre spectral sequence The transgression The path-loop fibre sequence 23

Homotopy Theory of Topological Spaces and Simplicial Sets

1 Replete topoi. X = Shv proét (X) X is locally weakly contractible (next lecture) X is replete. D(X ) is left complete. K D(X ) we have R lim

arxiv: v1 [math.ct] 10 Jul 2016

Internal Universes in Models of Homotopy Type Theory. Daniel R. Licata Ian Orton Andrew M. Pitts Bas Spitters

ENRICHED MODEL CATEGORIES AND PRESHEAF CATEGORIES

HOMOTOPY LIMITS IN TYPE THEORY JEREMY AVIGAD, KRZYSZTOF KAPULKIN, AND PETER LEFANU LUMSDAINE

QUILLEN MODEL STRUCTURES FOR RELATIVE HOMOLOGICAL ALGEBRA

arxiv: v1 [math.ct] 8 Apr 2019

THE CELLULARIZATION PRINCIPLE FOR QUILLEN ADJUNCTIONS

INDUCTIVE PRESENTATIONS OF GENERALIZED REEDY CATEGORIES. Contents. 1. The algebraic perspective on Reedy categories

MODEL-CATEGORIES OF COALGEBRAS OVER OPERADS

Cartesian Cubical Type Theory

The positive complete model structure and why we need it

PART IV.2. FORMAL MODULI

Homotopy type theory

Amalgamable diagram shapes

PARAMETRIZED HIGHER CATEGORY THEORY

Higher Inductive Types: The circle and friends, axiomatically

Higher Inductive types

Grothendieck duality for affine M 0 -schemes.

Combinatorial Models for M (Lecture 10)

Derived Algebraic Geometry IX: Closed Immersions

FLAGGED HIGHER CATEGORIES

ON -TOPOI JACOB LURIE

arxiv: v6 [math.ct] 2 Jul 2015


Galois descent criteria

An Introduction to Model Categories

FUNCTORS BETWEEN REEDY MODEL CATEGORIES OF DIAGRAMS

ENRICHED MODEL CATEGORIES IN EQUIVARIANT CONTEXTS

FINITE SPECTRA CARY MALKIEWICH

A Fibrational View of Geometric Morphisms

Moduli Problems for Structured Ring Spectra

Algebraic Models for Homotopy Types III Algebraic Models in p-adic Homotopy Theory

Axioms for Modelling Cubical Type Theory in a Topos

Transcription:

Categorical models of homotopy type theory Michael Shulman 12 April 2012

Outline 1 Homotopy type theory in model categories 2 The universal Kan fibration 3 Models in (, 1)-toposes

Homotopy type theory in higher categories Recall: homotopy type theory (, 1)-categories, + types products, coproducts equality types (x = y) diagonals types local cartesian closure univalent universe Type object classifier

Homotopy type theory in higher categories Recall: homotopy type theory (, 1)-categories, + types products, coproducts equality types (x = y) diagonals types local cartesian closure univalent universe Type object classifier

Two kinds of equality Problem Type theory is stricter than (, 1)-categories.

Two kinds of equality Problem Type theory is stricter than (, 1)-categories. In type theory, we have two kinds of equality : 1 Equality witnessed by inhabitants of equality types (= paths). 2 Computational equality: (λx.b)(a) evaluates to b[a/x].

Two kinds of equality Problem Type theory is stricter than (, 1)-categories. In type theory, we have two kinds of equality : 1 Equality witnessed by inhabitants of equality types (= paths). 2 Computational equality: (λx.b)(a) evaluates to b[a/x]. These play different roles: type checking depends on computational equality. if a evaluates to b, and c : C(a), then also c : C(b). In particular, if a evaluates to b, then refl b : (a = b). if p : (a = b) and c : C(a), then only transport(p, c): C(b).

Two kinds of equality But computational equality is also stricter. Example Composition is computationally strictly associative. g f := λx.g(f (x))

Two kinds of equality But computational equality is also stricter. Example Composition is computationally strictly associative. g f := λx.g(f (x)) ( (g ) ) h (g f ) = λx.h f (x)

Two kinds of equality But computational equality is also stricter. Example Composition is computationally strictly associative. g f := λx.g(f (x)) ( (λx.g(f ) ) h (g f ) = λx.h (x)) (x)

Two kinds of equality But computational equality is also stricter. Example Composition is computationally strictly associative. g f := λx.g(f (x)) ( (λx.g(f ) ) h (g f ) = λx.h (x)) (x) λx.h(g(f (x)))

Two kinds of equality But computational equality is also stricter. Example Composition is computationally strictly associative. g f := λx.g(f (x)) ( (λx.g(f ) ) h (g f ) = λx.h (x)) (x) λx.h(g(f (x))) ( ) (h g) f = λx. h g (f (x))

Two kinds of equality But computational equality is also stricter. Example Composition is computationally strictly associative. g f := λx.g(f (x)) ( (λx.g(f ) ) h (g f ) = λx.h (x)) (x) λx.h(g(f (x))) ( ) (h g) f = λx. λy.h(g(y)) (f (x))

Two kinds of equality But computational equality is also stricter. Example Composition is computationally strictly associative. g f := λx.g(f (x)) ( (λx.g(f ) ) h (g f ) = λx.h (x)) (x) λx.h(g(f (x))) ( ) (h g) f = λx. λy.h(g(y)) (f (x)) λx.h(g(f (x)))

Two kinds of equality But computational equality is also stricter. Example Composition is computationally strictly associative. g f := λx.g(f (x)) ( (λx.g(f ) ) h (g f ) = λx.h (x)) (x) λx.h(g(f (x))) ( ) (h g) f = λx. λy.h(g(y)) (f (x)) λx.h(g(f (x))) This is the sort of issue that homotopy theorists are intimately familiar with! We need a model for (, 1)-categories with (at least) a strictly associative composition law.

Display map categories Forget everything you know about homotopy theory; let s see how the type theorists come at it. Definition A display map category is a category with A terminal object. A subclass of its morphisms called the display maps, denoted P A or P A. Any pullback of a display map exists and is a display map.

Display map categories Forget everything you know about homotopy theory; let s see how the type theorists come at it. Definition A display map category is a category with A terminal object. A subclass of its morphisms called the display maps, denoted P A or P A. Any pullback of a display map exists and is a display map. A display map P A is a type dependent on A. A display map A 1 is a plain type (dependent on nothing). Pullback is substitution.

Dependent sums of display maps (x : A) (B(x) : Type) If the types B(x) are the fibers of B A, their dependent sum B(x) should be the object B. x : A B (x : A) (B(x) : Type) A 1 ( x : A B(x) : Type ) B 1

Dependent sums in context More generally: C (x : A), (y : B(x)) (C(x, y) : Type) B A (x : A) ( y : B(x) C(x, y) : Type ) C A

Dependent sums in context More generally: C (x : A), (y : B(x)) (C(x, y) : Type) B A (x : A) ( y : B(x) C(x, y) : Type ) C A Dependent sums display maps compose

Aside: adjoints to pullback In a category C, if pullbacks along f : A B exist, then the functor f : C /B C /A has a left adjoint Σ f given by composition with f.

Aside: adjoints to pullback In a category C, if pullbacks along f : A B exist, then the functor f : C /B C /A has a left adjoint Σ f given by composition with f. If f is a display map and display maps compose, then Σ f restricts to a functor (C /A) disp (C /B) disp implementing dependent sums.

Aside: adjoints to pullback In a category C, if pullbacks along f : A B exist, then the functor f : C /B C /A has a left adjoint Σ f given by composition with f. If f is a display map and display maps compose, then Σ f restricts to a functor implementing dependent sums. (C /A) disp (C /B) disp A right adjoint to f, if one exists, is an object of sections. C is locally cartesian closed iff all such right adjoints Π f exist.

Dependent products of display maps (x : A), (y : B(x)) (C(x, y) : Type) C B A (x : A) ( y : B(x) C(x, y) : Type ) Π B C B A

Dependent products of display maps (x : A), (y : B(x)) (C(x, y) : Type) C B A (x : A) ( y : B(x) C(x, y) : Type ) Π B C B A Dependent products display maps exponentiate

Identity types for display maps The dependent identity type (x : A), (y : A) ((x = y) : Type) must be a display map Id A A A

Identity types for display maps The reflexivity constructor must be a section (x : A) (refl(x) : (x = x)) Id A Id A A A A

Identity types for display maps The reflexivity constructor must be a section (x : A) (refl(x) : (x = x)) Id A Id A or equivalently a lifting A A A Id A refl A A A

Identity types for display maps The eliminator says given a dependent type with a section refl C A refl C Id A there exists a compatible section C Id A

Identity types for display maps The eliminator says given a dependent type with a section refl C A refl C Id A there exists a compatible section C Id A In other words, we have the lifting property refl A Id A C Id A

Identity types for display maps In fact, refl has the left lifting property w.r.t. all display maps. refl A Id A f C B

Identity types for display maps In fact, refl has the left lifting property w.r.t. all display maps. refl A f C Id A Id A f C B

Identity types for display maps In fact, refl has the left lifting property w.r.t. all display maps. refl A Id A f C Id A f C B

Identity types for display maps In fact, refl has the left lifting property w.r.t. all display maps. refl A Id A f C Id A f C B Conclusion Identity types factor : A A A as A refl q Id A A A where q is a display map and refl lifts against all display maps.

Weak factorization systems Definition We say j f if any commutative square j X Y B A f admits a (non-unique) diagonal filler.

Weak factorization systems Definition We say j f if any commutative square j X Y B A f admits a (non-unique) diagonal filler. J = { f j f j J } F = { j j f f F }

Weak factorization systems Definition We say j f if any commutative square j X Y B A f admits a (non-unique) diagonal filler. J = { f j f j J } F = { j j f f F } Definition A weak factorization system in a category is (J, F) such that 1 J = F and F = J. 2 Every morphism factors as f j for some f F and j J.

General factorizations Theorem (Gambino Garner) In a display map category that models identity types, any morphism g : A B factors as A j Ng f B where f is a display map, and j lifts against all display maps.

General factorizations Theorem (Gambino Garner) In a display map category that models identity types, any morphism g : A B factors as A j Ng f B where f is a display map, and j lifts against all display maps. (y : B) Ng(y) := hfiber(g, y) := x : A(g(x) = y) is the type-theoretic mapping path space.

The identity type wfs Corollary (Gambino-Garner) In a type theory with identity types, ( (display maps), ( (display maps)) ) is a weak factorization system.

The identity type wfs Corollary (Gambino-Garner) In a type theory with identity types, ( (display maps), ( (display maps)) ) is a weak factorization system. This behaves very much like (acyclic cofibrations, fibrations): Dependent types are like fibrations (recall transport ). Every map in (display maps) is an equivalence; in fact, the inclusion of a deformation retract.

Modeling identity types Conversely: Theorem (Awodey Warren,Garner van den Berg) In a display map category, if ( (display maps), ( (display maps)) ) is a pullback-stable weak factorization system, then the category (almost ) models identity types. identity types weak factorization systems

Model categories Definition (Quillen) A model category is a category C with limits and colimits and three classes of maps: C = cofibrations F = fibrations W = weak equivalences such that 1 W has the 2-out-of-3 property. 2 (C W, F) and (C, F W) are weak factorization systems.

Type-theoretic model categories Corollary Let M be a model category such that 1 M (as a category) is locally cartesian closed. 2 M is right proper. 3 The cofibrations are the monomorphisms. Then M (almost ) models type theory with dependent sums, dependent products, and identity types.

Type-theoretic model categories Corollary Let M be a model category such that 1 M (as a category) is locally cartesian closed. 2 M is right proper. 3 The cofibrations are the monomorphisms. Then M (almost ) models type theory with dependent sums, dependent products, and identity types. Homotopy theory (homotopy type) theory Type theory

Type-theoretic model categories Corollary Let M be a model category such that 1 M (as a category) is locally cartesian closed. 2 M is right proper. 3 The cofibrations are the monomorphisms. Then M (almost ) models type theory with dependent sums, dependent products, and identity types. Homotopy theory (homotopy type) theory Type theory Examples Simplicial sets with the Quillen model structure.

Type-theoretic model categories Corollary Let M be a model category such that 1 M (as a category) is locally cartesian closed. 2 M is right proper. 3 The cofibrations are the monomorphisms. Then M (almost ) models type theory with dependent sums, dependent products, and identity types. Homotopy theory (homotopy type) theory Type theory Examples Simplicial sets with the Quillen model structure. Any injective model structure on simplicial presheaves.

Homotopy type theory in categories (x : A) p : isprop(b(x)) (x : A), (u : B(x)), (v : B(x)) (p u,v : (u = v)) The path object P A B has a section in M/A Any two maps into B are homotopic over A

Homotopy type theory in categories (x : A) p : isprop(b(x)) (x : A), (u : B(x)), (v : B(x)) (p u,v : (u = v)) The path object P A B has a section in M/A Any two maps into B are homotopic over A (x : A) p : iscontr(b(x)) (x : A) p : isprop(b(x)) B(x) Any two maps into B are homotopic over A and B A has a section B A is an acyclic fibration

Homotopy type theory in categories For f : A B, p : isequiv(f ) iscontr(hfiber(f, y)) y : B (y : B) iscontr(hfiber(f, y)) hfiber(f ) A is an acyclic fibration f is a (weak) equivalence (Recall hfiber is the factorization A Nf B of f.)

Homotopy type theory in categories For f : A B, p : isequiv(f ) y : B iscontr(hfiber(f, y)) (y : B) iscontr(hfiber(f, y)) hfiber(f ) A is an acyclic fibration f is a (weak) equivalence (Recall hfiber is the factorization A Nf B of f.) Conclusion Any theorem about equivalences that we can prove in type theory yields a conclusion about weak equivalences in appropriate model categories.

Coherence Another Problem Type theory is even stricter than 1-categories! Recall that substitution is pullback. f g A g P P A f B g C a: A P(g(f (a))) b : B P(g(b)) c : C P(c)

Coherence Another Problem Type theory is even stricter than 1-categories! Recall that substitution is pullback. (g f ) A P A a: A P((g f )(a)) g f C c : C P(c)

Coherence Another Problem Type theory is even stricter than 1-categories! Recall that substitution is pullback. (g f ) A P A a: A P((λx.g(f (x)))(a)) g f C c : C P(c)

Coherence Another Problem Type theory is even stricter than 1-categories! Recall that substitution is pullback. (g f ) A P A a: A P(g(f (a))) g f C c : C P(c)

Coherence Another Problem Type theory is even stricter than 1-categories! Recall that substitution is pullback. (g f ) A P A a: A P(g(f (a))) g f C c : C P(c) But, of course, f g P is only isomorphic to (g f ) P.

Coherence with a universe There are several resolutions; perhaps the cleanest is: Solution (Voevodsky) Represent dependent types by their classifying maps into a universe object. Now substitution is composition, which is strictly associative (in our model category): A f B g C P U A g f C P U We needed a universe object anyway, to model the type Type and prove univalence.

Coherence with a universe There are several resolutions; perhaps the cleanest is: Solution (Voevodsky) Represent dependent types by their classifying maps into a universe object. Now substitution is composition, which is strictly associative (in our model category): A f B g C P U A g f C P U We needed a universe object anyway, to model the type Type and prove univalence. New problem Need very strict models for universe objects.

Outline 1 Homotopy type theory in model categories 2 The universal Kan fibration 3 Models in (, 1)-toposes

Representing fibrations (Following Kapulkin Lumsdaine Voevodsky) Goal A universe object in simplicial sets giving coherence and univalence.

Representing fibrations (Following Kapulkin Lumsdaine Voevodsky) Goal A universe object in simplicial sets giving coherence and univalence. Simplicial sets are a presheaf category, so there is a standard trick to build representing objects. U n = Hom( n, U) {fibrations over n }

Representing fibrations (Following Kapulkin Lumsdaine Voevodsky) Goal A universe object in simplicial sets giving coherence and univalence. Simplicial sets are a presheaf category, so there is a standard trick to build representing objects. U n = Hom( n, U) {fibrations over n } But n {fibrations over n } is only a pseudofunctor; we need to rigidify it.

Well-ordered fibrations A technical device (Voevodsky) A well-ordered Kan fibration is a Kan fibration p : E B together with, for every x B n, a well-ordering on p 1 (x) E n.

Well-ordered fibrations A technical device (Voevodsky) A well-ordered Kan fibration is a Kan fibration p : E B together with, for every x B n, a well-ordering on p 1 (x) E n. Two well-ordered Kan fibrations are isomorphic in at most one way which preserves the orders.

Well-ordered fibrations A technical device (Voevodsky) A well-ordered Kan fibration is a Kan fibration p : E B together with, for every x B n, a well-ordering on p 1 (x) E n. Two well-ordered Kan fibrations are isomorphic in at most one way which preserves the orders. Definition U n := { X n a well-ordered fibration }/ ordered =

Well-ordered fibrations A technical device (Voevodsky) A well-ordered Kan fibration is a Kan fibration p : E B together with, for every x B n, a well-ordering on p 1 (x) E n. Two well-ordered Kan fibrations are isomorphic in at most one way which preserves the orders. Definition Ũ n := U n := { X n a well-ordered fibration }/ ordered = { }/ (X, x) X n well-ordered fibration, x X n ordered =

Well-ordered fibrations A technical device (Voevodsky) A well-ordered Kan fibration is a Kan fibration p : E B together with, for every x B n, a well-ordering on p 1 (x) E n. Two well-ordered Kan fibrations are isomorphic in at most one way which preserves the orders. Definition Ũ n := U n := { X n a well-ordered fibration }/ ordered = { }/ (X, x) X n well-ordered fibration, x X n (with some size restriction, to make them sets). ordered =

The universal Kan fibration Theorem The forgetful map Ũ U is a Kan fibration. Proof. A map E B is a Kan fibration if and only if every pullback b E n b E B is such, since the horns Λ n k n have codomain n.

The universal Kan fibration Theorem The forgetful map Ũ U is a Kan fibration. Proof. A map E B is a Kan fibration if and only if every pullback b E n b E B is such, since the horns Λ n k n have codomain n. Thus, of course, every pullback of Ũ U is a Kan fibration.

The universal Kan fibration Theorem Every (small) Kan fibration E B is some pullback of Ũ U: E B Ũ U Proof. Choose a well-ordering on each fiber, and map x B n to the isomorphism class of the well-ordered fibration b (E) n.

The universal Kan fibration Theorem Every (small) Kan fibration E B is some pullback of Ũ U: E B Ũ U Proof. Choose a well-ordering on each fiber, and map x B n to the isomorphism class of the well-ordered fibration b (E) n. It is essential that we have actual pullbacks here, not just homotopy pullbacks.

Type theory in the universe Let the size-bound for U be inaccessible (a Grothendieck universe). Then small fibrations are closed under all categorical constructions.

Type theory in the universe Let the size-bound for U be inaccessible (a Grothendieck universe). Then small fibrations are closed under all categorical constructions. Now we can interpret type theory with coherence, using morphisms into U for dependent types. Example A context (x : A), (y : B(x)), (z : C(x, y)) becomes a sequence of fibrations together with classifying maps: C B A 1 [C] [B] [A] Ũ U Ũ U Ũ U in which each trapezoid is a pullback.

Strict cartesian products Every type-theoretic operation can be done once over U, then implemented by composition. Example (Cartesian product) Pull Ũ back to U U along the two projections π 1, π 2.

Strict cartesian products Every type-theoretic operation can be done once over U, then implemented by composition. Example (Cartesian product) Pull Ũ back to U U along the two projections π 1, π 2. Their fiber product over U U admits a classifying map: (π 1Ũ) U U (π 2Ũ) Ũ U U [ ] U

Strict cartesian products Every type-theoretic operation can be done once over U, then implemented by composition. Example (Cartesian product) Pull Ũ back to U U along the two projections π 1, π 2. Their fiber product over U U admits a classifying map: (π 1Ũ) U U (π 2Ũ) Ũ U U [ ] U Define the product of [A]: X U and [B]: X U to be This has strict substitution. X ([A],[B]) U U [ ] U

Nested universes Problem So far the object U lives outside the type theory. We want it inside, giving a universe type Type and univalence.

Nested universes Problem So far the object U lives outside the type theory. We want it inside, giving a universe type Type and univalence. Solution Let U be a bigger universe. If U is U -small and fibrant, then it has a classifying map: U Ũ 1 u U and the type theory defined using U has a universe type u.

U is fibrant Theorem U is fibrant.

U is fibrant Theorem U is fibrant. Outline of proof. Λ n k f U j n?

U is fibrant Theorem U is fibrant. Outline of proof. Λ n k f U j n? With hard work, we can extend f Ũ to a fibration over n : f Ũ Λ n k j P n

U is fibrant Theorem U is fibrant. Outline of proof. Λ n k f U j n? With hard work, we can extend f Ũ to a fibration over n : f Ũ Λ n k j P n and extend the well-ordering of f Ũ to P, yielding g : n U with gj = f (and g Ũ = P).

U is fibrant Theorem U is fibrant. Outline of proof. Λ n k f U j n? With hard work, we can extend f Ũ to a fibration over n : f Ũ Λ n k j P n and extend the well-ordering of f Ũ to P, yielding g : n U with gj = f (and g Ũ = P).

Extending fibrations Lemma Any fibration P Λ n k is the pullback of some fibration over n.

Extending fibrations Lemma Any fibration P Λ n k is the pullback of some fibration over n. Proof. Let P P be a minimal subfibration.

Extending fibrations Lemma Any fibration P Λ n k is the pullback of some fibration over n. Proof. Let P P be a minimal subfibration. There is a retraction P P that is an acyclic fibration.

Extending fibrations Lemma Any fibration P Λ n k is the pullback of some fibration over n. Proof. Let P P be a minimal subfibration. There is a retraction P P that is an acyclic fibration. Since Λ n k is contractible, the minimal fibration P Λ n k is isomorphic to a trivial bundle Λ n k F Λn k.

Extending fibrations Lemma Any fibration P Λ n k is the pullback of some fibration over n. Proof. Let P P be a minimal subfibration. There is a retraction P P that is an acyclic fibration. Since Λ n k is contractible, the minimal fibration P Λ n k is isomorphic to a trivial bundle Λ n k F Λn k. P P = Λ n k F Λ n k j n

Extending fibrations Lemma Any fibration P Λ n k is the pullback of some fibration over n. Proof. Let P P be a minimal subfibration. There is a retraction P P that is an acyclic fibration. Since Λ n k is contractible, the minimal fibration P Λ n k is isomorphic to a trivial bundle Λ n k F Λn k. P P = Λ n k F j F n F Λ n k j n

Extending fibrations Lemma Any fibration P Λ n k is the pullback of some fibration over n. Proof. Let P P be a minimal subfibration. There is a retraction P P that is an acyclic fibration. Since Λ n k is contractible, the minimal fibration P Λ n k is isomorphic to a trivial bundle Λ n k F Λn k. P Π j F P P = Λ n k F j F n F Λ n k j n

Univalence We want to show that PU Eq(U) is an equivalence: PU? Eq(U) U U

Univalence We want to show that PU Eq(U) is an equivalence: U PU? Eq(U) U U It suffices to show: 1 The composite U Eq(U) is an equivalence.

Univalence We want to show that PU Eq(U) is an equivalence: U PU? Eq(U) id U U U π 2? It suffices to show: 1 The composite U Eq(U) is an equivalence. 2 The projection Eq(U) U is an equivalence.

Univalence We want to show that PU Eq(U) is an equivalence: U PU? Eq(U) id U U U π 2? It suffices to show: 1 The composite U Eq(U) is an equivalence. 2 The projection Eq(U) U is an equivalence. 3 The projection Eq(U) U is an acyclic fibration.

Univalence By representability, a commutative square n Eq(U) i n U corresponds to a diagram E 1 E 2 E 2 n with E 1 E 2 an equivalence. i n

Univalence By representability, a commutative square with a lift n Eq(U) i n U corresponds to a diagram E 1 E 1 E 2 E 2 n i n with E 1 E 2 and E 1 E 2 equivalences.

Univalence E 1 E 1 Π i (E 1 ) E 2 e 2 E 2 Π i (E 2 ) n i n

Univalence E 1 E 1 Π i (E 1 ) E 2 e 2 E 2 Π i (E 2 ) n i n By factorization, consider separately the cases when E 1 E 2 is (1) an acyclic fibration or (2) an acyclic cofibration.

Univalence E 1 E 1 Π i (E 1 ) E 2 e 2 E 2 Π i (E 2 ) n i n By factorization, consider separately the cases when E 1 E 2 is (1) an acyclic fibration or (2) an acyclic cofibration. (1) E 1 E 2 is an acyclic fibration (Π i preserves such).

Univalence E 1 E 1 Π i (E 1 ) E 2 e 2 E 2 Π i (E 2 ) n i n By factorization, consider separately the cases when E 1 E 2 is (1) an acyclic fibration or (2) an acyclic cofibration. (1) E 1 E 2 is an acyclic fibration (Π i preserves such). (2) E 1 is a deformation retract of E 2.

Outline 1 Homotopy type theory in model categories 2 The universal Kan fibration 3 Models in (, 1)-toposes

(, 1)-toposes Definition An (, 1)-topos is an (, 1)-category that is a left-exact localization of an (, 1)-presheaf category. Examples -groupoids (plays the role of the 1-topos Set) Parametrized homotopy theory over any space X G-equivariant homotopy theory for any group G -sheaves/stacks on any space Smooth -groupoids (or algebraic etc.)

Univalence in categories Definition (Rezk) An object classifier in an (, 1)-category C is a morphism Ũ U such that pullback B Ũ induces an equivalence of -groupoids Hom(A, U) Core(C/A) small A U ( Core is the maximal sub- -groupoid.)

(, 1)-toposes Theorem (Rezk) An (, 1)-category C is an (, 1)-topos if and only if 1 C is locally presentable. 2 C is locally cartesian closed. 3 κ-compact objects have object classifiers for κ 0.

(, 1)-toposes Theorem (Rezk) An (, 1)-category C is an (, 1)-topos if and only if 1 C is locally presentable. 2 C is locally cartesian closed. 3 κ-compact objects have object classifiers for κ 0. Corollary If a combinatorial model category M interprets dependent type theory as before (i.e. it is locally cartesian closed, right proper, and the cofibrations are the monomorphisms), and contains universes for κ-compact objects that satisfy the univalence axiom, then the (, 1)-category that it presents is an (, 1)-topos.

(, 1)-toposes Conjecture Every (, 1)-topos can be presented by a model category which interprets dependent type theory with the univalence axiom. Homotopy type theory is the internal logic of (, 1)-toposes.

(, 1)-toposes Conjecture Every (, 1)-topos can be presented by a model category which interprets dependent type theory with the univalence axiom. Homotopy type theory is the internal logic of (, 1)-toposes. If this is true, then anything we prove in homotopy type theory (which we can also verify with a computer) will automatically be true internally to any (, 1)-topos. The constructive core of homotopy theory should be provable in this way, in a uniform way for all homotopy theories.

Status of the conjecture Gpd (, 1)-presheaves (, 1)-toposes

Status of the conjecture Gpd (, 1)-presheaves (, 1)-toposes

Status of the conjecture Gpd (, 1)-presheaves (, 1)-toposes

Status of the conjecture Gpd (, 1)-presheaves (, 1)-toposes inverse (, 1)-presheaves

Status of the conjecture Gpd (, 1)-presheaves (, 1)-toposes? inverse (, 1)-presheaves