Built with Alectryon, running Coq+SerAPI v8.18.0+0.18.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use instead of Ctrl.
(* -*- mode: coq; mode: visual-line -*- *)
[Loading ML file number_string_notation_plugin.cmxs (using legacy method) ... done]
Require Import Equiv.BiInv Extensions HProp HFiber NullHomotopy Limits.Pullback. Require Import PathAny. Require Import Colimits.Pushout. Local Open Scope nat_scope. Local Open Scope path_scope. (** * Reflective Subuniverses *) (** ** References *) (** Reflective subuniverses (and modalities) are studied in the following papers, which we will refer to below by their abbreviations: - The Book: The Homotopy Type Theory Book, chapter 7. Bare references to "Theorem 7.x.x" are always to the Book. - RSS: Rijke, Spitters, and Shulman, "Modalities in homotopy type theory", https://arxiv.org/abs/1706.07526. - CORS: Christensen, Opie, Rijke, and Scoccola, "Localization in Homotopy Type Theory", https://arxiv.org/abs/1807.04155. *) (** ** Definitions *) (** *** Subuniverses *) Record Subuniverse@{i} := { In_internal : Type@{i} -> Type@{i} ; hprop_inO_internal : Funext -> forall (T : Type@{i}), IsHProp (In_internal T) ; inO_equiv_inO_internal : forall (T U : Type@{i}) (T_inO : In_internal T) (f : T -> U) {feq : IsEquiv f}, In_internal U ; }. (** Work around Coq bug that fields of records can't be typeclasses. *) Class In (O : Subuniverse) (T : Type) := in_internal : In_internal O T. (** Being in the subuniverse is a mere predicate (by hypothesis). We include funext in the hypotheses of hprop_inO so that it doesn't have to be assumed in all definitions of (reflective) subuniverses, since in most examples it is required for this and this only. Here we redefine it using the replaced [In]. *) Global Instance hprop_inO `{Funext} (O : Subuniverse) (T : Type) : IsHProp (In O T) := @hprop_inO_internal _ _ T. (** We assumed repleteness of the subuniverse in the definition. Of course, with univalence this would be automatic, but we include it as a hypothesis since most of the theory of reflective subuniverses and modalities doesn't need univalence, and most or all examples can be shown to be replete without using univalence. Here we redefine it using the replaced [In]. *) Definition inO_equiv_inO {O : Subuniverse} (T : Type) {U : Type} `{T_inO : In O T} (f : T -> U) `{IsEquiv T U f} : In O U := @inO_equiv_inO_internal O T U T_inO f _. Definition inO_equiv_inO' {O : Subuniverse} (T : Type) {U : Type} `{In O T} (f : T <~> U) : In O U := inO_equiv_inO T f. Definition iff_inO_equiv (O : Subuniverse) {T : Type} {U : Type} (f : T <~> U) : In O T <-> In O U := (fun H => inO_equiv_inO' _ f, fun H => inO_equiv_inO' _ f^-1). Definition equiv_inO_equiv `{Funext} (O : Subuniverse) {T : Type} {U : Type} (f : T <~> U) : In O T <~> In O U := equiv_iff_hprop_uncurried (iff_inO_equiv _ f). (** The universe of types in the subuniverse *) Definition Type_@{i j} (O : Subuniverse@{i}) : Type@{j} := @sig@{j i} Type@{i} (fun (T : Type@{i}) => In O T). Coercion TypeO_pr1 O (T : Type_ O) := @pr1 Type (In O) T. (** The second component of [TypeO] is unique. *) Definition path_TypeO@{i j} {fs : Funext} O (T T' : Type_@{i j} O) (p : T.1 = T'.1) : T = T' := path_sigma_hprop@{j i j} T T' p. Definition equiv_path_TypeO@{i j} {fs : Funext} O (T T' : Type_@{i j} O) : (paths@{j} T.1 T'.1) <~> (T = T') := equiv_path_sigma_hprop@{j i j} T T'. (** Types in [TypeO] are always in [O]. *) Global Instance inO_TypeO {O : Subuniverse} (A : Type_ O) : In O A := A.2. (** ** Properties of Subuniverses *) (** A map is O-local if all its fibers are. *) Class MapIn (O : Subuniverse) {A B : Type} (f : A -> B) := inO_hfiber_ino_map : forall (b:B), In O (hfiber f b). Global Existing Instance inO_hfiber_ino_map. Section Subuniverse. Context (O : Subuniverse). (** Being a local map is an hprop *)
O: Subuniverse
H: Funext
A, B: Type
f: A -> B

IsHProp (MapIn O f)
O: Subuniverse
H: Funext
A, B: Type
f: A -> B

IsHProp (MapIn O f)
apply istrunc_forall. Defined. (** Anything homotopic to a local map is local. *)
O: Subuniverse
A, B: Type
f, g: A -> B
p: f == g
H: MapIn O f

MapIn O g
O: Subuniverse
A, B: Type
f, g: A -> B
p: f == g
H: MapIn O f

MapIn O g
O: Subuniverse
A, B: Type
f, g: A -> B
p: f == g
H: MapIn O f
b: B

In O (hfiber g b)
exact (inO_equiv_inO (hfiber f b) (equiv_hfiber_homotopic f g p b)). Defined. (** The projection from a family of local types is local. *)
O: Subuniverse
A: Type
B: A -> Type
H: forall a : A, In O (B a)

MapIn O pr1
O: Subuniverse
A: Type
B: A -> Type
H: forall a : A, In O (B a)

MapIn O pr1
O: Subuniverse
A: Type
B: A -> Type
H: forall a : A, In O (B a)
a: A

In O (hfiber pr1 a)
exact (inO_equiv_inO (B a) (hfiber_fibration a B)). Defined. (** A family of types is local if and only if the associated projection map is local. *)
O: Subuniverse
A: Type
B: A -> Type

(forall a : A, In O (B a)) <-> MapIn O pr1
O: Subuniverse
A: Type
B: A -> Type

(forall a : A, In O (B a)) <-> MapIn O pr1
O: Subuniverse
A: Type
B: A -> Type

(forall a : A, In O (B a)) -> MapIn O pr1
O: Subuniverse
A: Type
B: A -> Type
MapIn O pr1 -> forall a : A, In O (B a)
O: Subuniverse
A: Type
B: A -> Type

(forall a : A, In O (B a)) -> MapIn O pr1
exact _. (* Uses the instance mapinO_pr1 above. *)
O: Subuniverse
A: Type
B: A -> Type

MapIn O pr1 -> forall a : A, In O (B a)
O: Subuniverse
A: Type
B: A -> Type
a: A
x: In O (hfiber pr1 (?Goal a))

In O (B a)
exact (inO_equiv_inO (hfiber pr1 a) (hfiber_fibration a B)^-1%equiv). Defined.
O: Subuniverse
H: Funext
A: Type
B: A -> Type

(forall a : A, In O (B a)) <~> MapIn O pr1
O: Subuniverse
H: Funext
A: Type
B: A -> Type

(forall a : A, In O (B a)) <~> MapIn O pr1
exact (equiv_iff_hprop_uncurried (iff_forall_inO_mapinO_pr1 B)). Defined. End Subuniverse. (** *** Reflections *) (** A pre-reflection is a map to a type in the subuniverse. *) Class PreReflects@{i} (O : Subuniverse@{i}) (T : Type@{i}) := { O_reflector : Type@{i} ; O_inO : In O O_reflector ; to : T -> O_reflector ; }. Arguments O_reflector O T {_}. Arguments to O T {_}. Arguments O_inO {O} T {_}. Global Existing Instance O_inO. (** It is a reflection if it has the requisite universal property. *) Class Reflects@{i} (O : Subuniverse@{i}) (T : Type@{i}) `{PreReflects@{i} O T} := { extendable_to_O : forall {Q : Type@{i}} {Q_inO : In O Q}, ooExtendableAlong (to O T) (fun _ => Q) }. Arguments extendable_to_O O {T _ _ Q Q_inO}. (** Here's a modified version that applies to types in possibly-smaller universes without collapsing those universes to [i]. *)
O: Subuniverse
T: Type
H: PreReflects O T
H0: Reflects O T
Q: Type
Q_inO: In O Q

ooExtendableAlong (to O T) (fun _ : O_reflector O T => Q)
O: Subuniverse
T: Type
H: PreReflects O T
H0: Reflects O T
Q: Type
Q_inO: In O Q

ooExtendableAlong (to O T) (fun _ : O_reflector O T => Q)
O: Subuniverse
T: Type
H: PreReflects O T
H0: Reflects O T
Q: Type
Q_inO: In O Q

ooExtendableAlong (to O T) (fun _ : O_reflector O T => Q)
rapply extendable_to_O. Defined. (** In particular, every type in the subuniverse automatically reflects into it. *)
O: Subuniverse
T: Type
H: In O T

PreReflects O T
O: Subuniverse
T: Type
H: In O T

PreReflects O T
O: Subuniverse
T: Type
H: In O T

Type
O: Subuniverse
T: Type
H: In O T
In O ?O_reflector
O: Subuniverse
T: Type
H: In O T
T -> ?O_reflector
O: Subuniverse
T: Type
H: In O T

Type
exact T.
O: Subuniverse
T: Type
H: In O T

In O T
assumption.
O: Subuniverse
T: Type
H: In O T

T -> T
exact idmap. Defined.
O: Subuniverse
T: Type
H: In O T

Reflects O T
O: Subuniverse
T: Type
H: In O T

Reflects O T
constructor; intros; rapply ooextendable_equiv. Defined. (** A reflective subuniverse is one for which every type reflects into it. *) Record ReflectiveSubuniverse@{i} := { rsu_subuniv : Subuniverse@{i} ; rsu_prereflects : forall (T : Type@{i}), PreReflects rsu_subuniv T ; rsu_reflects : forall (T : Type@{i}), Reflects rsu_subuniv T ; }. Coercion rsu_subuniv : ReflectiveSubuniverse >-> Subuniverse. Global Existing Instance rsu_prereflects. Global Existing Instance rsu_reflects. (** We allow the name of a subuniverse or modality to be used as the name of its reflector. This means that when defining a particular example, you should generally put the parametrizing family in a wrapper, so that you can notate the subuniverse as parametrized by, rather than identical to, its parameter. See Modality.v, Truncations.v, and Localization.v for examples. *) Definition rsu_reflector (O : ReflectiveSubuniverse) (T : Type) : Type := O_reflector O T. Coercion rsu_reflector : ReflectiveSubuniverse >-> Funclass. (** *** Recursion principles *) (** We now extract the recursion principle and the restricted induction principles for paths. *) Section ORecursion. Context {O : Subuniverse} {P Q : Type} {Q_inO : In O Q} `{Reflects O P}. Definition O_rec (f : P -> Q) : O_reflector O P -> Q := (fst (extendable_to_O O 1%nat) f).1. Definition O_rec_beta (f : P -> Q) (x : P) : O_rec f (to O P x) = f x := (fst (extendable_to_O O 1%nat) f).2 x. Definition O_indpaths (g h : O_reflector O P -> Q) (p : g o to O P == h o to O P) : g == h := (fst (snd (extendable_to_O O 2) g h) p).1. Definition O_indpaths_beta (g h : O_reflector O P -> Q) (p : g o (to O P) == h o (to O P)) (x : P) : O_indpaths g h p (to O P x) = p x := (fst (snd (extendable_to_O O 2) g h) p).2 x. Definition O_ind2paths {g h : O_reflector O P -> Q} (p q : g == h) (r : p oD (to O P) == q oD (to O P)) : p == q := (fst (snd (snd (extendable_to_O O 3) g h) p q) r).1. Definition O_ind2paths_beta {g h : O_reflector O P -> Q} (p q : g == h) (r : p oD (to O P) == q oD (to O P)) (x : P) : O_ind2paths p q r (to O P x) = r x := (fst (snd (snd (extendable_to_O O 3) g h) p q) r).2 x. (** Clearly we can continue indefinitely as needed. *) End ORecursion. (* We never want to see [extendable_to_O]. The [!x] allows [cbn] to unfold these when passed a constructor, such as [tr x]. This, for example, means that [O_rec (O:=Tr n) f (tr x)] will compute to [f x] and [Trunc_functor n f (tr x)] will compute to [tr (f x)]. *) Arguments O_rec {O} {P Q}%type_scope {Q_inO H H0} f%function_scope !x. Arguments O_rec_beta {O} {P Q}%type_scope {Q_inO H H0} f%function_scope !x. Arguments O_indpaths {O} {P Q}%type_scope {Q_inO H H0} (g h)%function_scope p !x. Arguments O_indpaths_beta {O} {P Q}%type_scope {Q_inO H H0} (g h)%function_scope p !x. Arguments O_ind2paths {O} {P Q}%type_scope {Q_inO H H0} {g h}%function_scope p q r !x. Arguments O_ind2paths_beta {O} {P Q}%type_scope {Q_inO H H0} {g h}%function_scope p q r !x. (** A tactic that generalizes [strip_truncations] to reflective subuniverses. [strip_truncations] introduces fewer universe variables, so tends to work better when removing truncations. [strip_modalities] in Modality.v also applies dependent elimination when [O] is a modality. *) Ltac strip_reflections := (** Search for hypotheses of type [O X] for some [O] such that the goal is [O]-local. *) progress repeat match goal with | [ T : _ |- _ ] => revert_opaque T; refine (@O_rec _ _ _ _ _ _ _) || refine (@O_indpaths _ _ _ _ _ _ _ _ _); (** Ensure that we didn't generate more than one subgoal, i.e. that the goal was appropriately local. *) []; intro T end. (** Given [Funext], we prove the definition of reflective subuniverse in the book. *) Global Instance isequiv_o_to_O `{Funext} (O : ReflectiveSubuniverse) (P Q : Type) `{In O Q} : IsEquiv (fun g : O P -> Q => g o to O P) := isequiv_ooextendable _ _ (extendable_to_O O). Definition equiv_o_to_O `{Funext} (O : ReflectiveSubuniverse) (P Q : Type) `{In O Q} : (O P -> Q) <~> (P -> Q) := Build_Equiv _ _ (fun g : O P -> Q => g o to O P) _. (** [isequiv_ooextendable] is defined in a way that makes [O_rec] definitionally equal to the inverse of [equiv_o_to_O]. *) Global Instance isequiv_O_rec_to_O `{Funext} (O : ReflectiveSubuniverse) (P Q : Type) `{In O Q} : IsEquiv (fun g : P -> Q => O_rec g) := (equiv_isequiv (equiv_o_to_O O P Q)^-1). (** ** Properties of Reflective Subuniverses *) (** We now prove a bunch of things about an arbitrary reflective subuniverse. *) Section Reflective_Subuniverse. Context (O : ReflectiveSubuniverse). (** Functoriality of [O_rec] homotopies *)
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g

O_rec f == O_rec g
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g

O_rec f == O_rec g
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

O_rec f (to O P x) = O_rec g (to O P x)
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

O_rec f (to O P x) = ?Goal
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P
?Goal = O_rec g (to O P x)
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

O_rec f (to O P x) = ?Goal
apply O_rec_beta.
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

f x = O_rec g (to O P x)
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

f x = O_rec g (to O P x)
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

f x = ?Goal
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P
?Goal = O_rec g (to O P x)
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

f x = ?Goal
exact (pi _).
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

g x = O_rec g (to O P x)
O: ReflectiveSubuniverse
P, Q: Type
H: In O Q
f, g: P -> Q
pi: f == g
x: P

g x = O_rec g (to O P x)
symmetry; apply O_rec_beta. } } Defined. (** If [T] is in the subuniverse, then [to O T] is an equivalence. *)
O: ReflectiveSubuniverse
T: Type
H: In O T

IsEquiv (to O T)
O: ReflectiveSubuniverse
T: Type
H: In O T

IsEquiv (to O T)
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T

IsEquiv (to O T)
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T

(fun x : O_reflector O T => to O T (g x)) == idmap
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T
(fun x : T => g (to O T x)) == idmap
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T

(fun x : O_reflector O T => to O T (g x)) == idmap
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T

(fun x : T => to O T (g (to O T x))) == (fun x : T => to O T x)
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T
x: T

to O T (g (to O T x)) = to O T x
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T
x: T

g (to O T x) = x
apply O_rec_beta.
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T

(fun x : T => g (to O T x)) == idmap
O: ReflectiveSubuniverse
T: Type
H: In O T
g:= O_rec idmap : O T -> T: O T -> T
x: T

g (to O T x) = x
apply O_rec_beta. Defined. Definition equiv_to_O (T : Type) `{In O T} : T <~> O T := Build_Equiv T (O T) (to O T) _. Section Functor. (** In this section, we see that [O] is a functor. *) Definition O_functor {A B : Type} (f : A -> B) : O A -> O B := O_rec (to O B o f). (** Naturality of [to O] *) Definition to_O_natural {A B : Type} (f : A -> B) : (O_functor f) o (to O A) == (to O B) o f := (O_rec_beta _). (** Functoriality on composition *)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C

O_functor (g o f) == O_functor g o O_functor f
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C

O_functor (g o f) == O_functor g o O_functor f
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
x: A

O_functor (fun x : A => g (f x)) (to O A x) = O_functor g (O_functor f (to O A x))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
x: A

to O C (g (f x)) = O_functor g (O_functor f (to O A x))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
x: A

to O C (g (f x)) = O_functor g (to O B (f x))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
x: A
O_functor g (to O B (f x)) = O_functor g (O_functor f (to O A x))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
x: A

to O C (g (f x)) = O_functor g (to O B (f x))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
x: A

O_functor g (to O B (f x)) = to O C (g (f x))
exact (to_O_natural g (f x)).
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
x: A

O_functor g (to O B (f x)) = O_functor g (O_functor f (to O A x))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
x: A

O_functor f (to O A x) = to O B (f x)
exact (to_O_natural f x). Defined. (** Functoriality on homotopies (2-functoriality) *)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g

O_functor f == O_functor g
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g

O_functor f == O_functor g
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

O_functor f (to O A x) = O_functor g (to O A x)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

to O B (f x) = O_functor g (to O A x)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

to O B (f x) = to O B (g x)
apply ap, pi. Defined. (** Functoriality for inverses of homotopies *)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g

O_functor_homotopy g f (fun x : A => (pi x)^) == (fun x : O A => (O_functor_homotopy f g pi x)^)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g

O_functor_homotopy g f (fun x : A => (pi x)^) == (fun x : O A => (O_functor_homotopy f g pi x)^)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

(O_functor_homotopy g f (fun x : A => (pi x)^) oD to O A) x = ((fun x : O A => (O_functor_homotopy f g pi x)^) oD to O A) x
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

O_indpaths (O_functor g) (O_functor f) (fun x : A => to_O_natural g x @ (ap (to O B) (pi x)^ @ (to_O_natural f x)^)) (to O A x) = (O_indpaths (O_functor f) (O_functor g) (fun x : A => to_O_natural f x @ (ap (to O B) (pi x) @ (to_O_natural g x)^)) (to O A x))^
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

(to_O_natural g x @ (ap (to O B) (pi x))^) @ (to_O_natural f x)^ = (to_O_natural g x @ (ap (to O B) (pi x))^) @ (to_O_natural f x)^
reflexivity. Qed. (** Hence functoriality on commutative squares *)
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2

O_functor f o O_functor pi1 == O_functor g o O_functor pi2
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2

O_functor f o O_functor pi1 == O_functor g o O_functor pi2
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X

O_functor f (O_functor pi1 x) = O_functor g (O_functor pi2 x)
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X

O_functor f (O_functor pi1 x) = O_functor (f o pi1) x
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X
O_functor (f o pi1) x = O_functor g (O_functor pi2 x)
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X

O_functor f (O_functor pi1 x) = O_functor (f o pi1) x
symmetry; rapply O_functor_compose.
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X

O_functor (f o pi1) x = O_functor g (O_functor pi2 x)
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X

O_functor (f o pi1) x = O_functor (g o pi2) x
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X
O_functor (g o pi2) x = O_functor g (O_functor pi2 x)
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X

O_functor (f o pi1) x = O_functor (g o pi2) x
apply O_functor_homotopy, comm.
O: ReflectiveSubuniverse
A, B, C, X: Type
pi1: X -> A
pi2: X -> B
f: A -> C
g: B -> C
comm: f o pi1 == g o pi2
x: O X

O_functor (g o pi2) x = O_functor g (O_functor pi2 x)
rapply O_functor_compose. Defined. (** Functoriality on identities *)
O: ReflectiveSubuniverse
A: Type

O_functor idmap == idmap
O: ReflectiveSubuniverse
A: Type

O_functor idmap == idmap
O: ReflectiveSubuniverse
A: Type
x: A

O_functor idmap (to O A x) = to O A x
apply O_rec_beta. Defined. (** 3-functoriality, as an example use of [O_ind2paths] *)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
p, q: f == g
r: p == q

O_functor_homotopy f g p == O_functor_homotopy f g q
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
p, q: f == g
r: p == q

O_functor_homotopy f g p == O_functor_homotopy f g q
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
p, q: f == g
r: p == q
x: A

(O_functor_homotopy f g p oD to O A) x = (O_functor_homotopy f g q oD to O A) x
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
p, q: f == g
r: p == q
x: A

O_indpaths (O_functor f) (O_functor g) (fun x : A => to_O_natural f x @ (ap (to O B) (p x) @ (to_O_natural g x)^)) (to O A x) = O_indpaths (O_functor f) (O_functor g) (fun x : A => to_O_natural f x @ (ap (to O B) (q x) @ (to_O_natural g x)^)) (to O A x)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
p, q: f == g
r: p == q
x: A

to_O_natural f x @ (ap (to O B) (p x) @ (to_O_natural g x)^) = to_O_natural f x @ (ap (to O B) (q x) @ (to_O_natural g x)^)
apply whiskerL, whiskerR, ap, r. (** Of course, if we wanted to prove 4-functoriality, we'd need to make this transparent. *) Qed. (** 2-naturality: Functoriality on homotopies is also natural *)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

O_functor_homotopy f g pi (to O A x) = (to_O_natural f x @ ap (to O B) (pi x)) @ (to_O_natural g x)^
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

O_functor_homotopy f g pi (to O A x) = (to_O_natural f x @ ap (to O B) (pi x)) @ (to_O_natural g x)^
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

O_indpaths (O_functor f) (O_functor g) (fun x : A => O_rec_beta (fun x0 : A => to O B (f x0)) x @ (ap (to O B) (pi x) @ (O_rec_beta (fun x0 : A => to O B (g x0)) x)^)) (to O A x) = (O_rec_beta (fun x : A => to O B (f x)) x @ ap (to O B) (pi x)) @ (O_rec_beta (fun x : A => to O B (g x)) x)^
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
pi: f == g
x: A

O_rec_beta (fun x : A => to O B (f x)) x @ (ap (to O B) (pi x) @ (O_rec_beta (fun x : A => to O B (g x)) x)^) = (O_rec_beta (fun x : A => to O B (f x)) x @ ap (to O B) (pi x)) @ (O_rec_beta (fun x : A => to O B (g x)) x)^
refine (concat_p_pp _ _ _). Defined. (** The pointed endofunctor ([O],[to O]) is well-pointed *)
O: ReflectiveSubuniverse
A: Type

O_functor (to O A) == to O (O_reflector O A)
O: ReflectiveSubuniverse
A: Type

O_functor (to O A) == to O (O_reflector O A)
O: ReflectiveSubuniverse
A: Type
x: A

O_functor (to O A) (to O A x) = to O (O_reflector O A) (to O A x)
apply to_O_natural. Defined. (** "Functoriality of naturality": the pseudonaturality axiom for composition *)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
a: A

ap (O_functor g) (to_O_natural f a) @ to_O_natural g (f a) = (O_functor_compose f g (to O A a))^ @ to_O_natural (g o f) a
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
a: A

ap (O_functor g) (to_O_natural f a) @ to_O_natural g (f a) = (O_functor_compose f g (to O A a))^ @ to_O_natural (g o f) a
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
a: A

ap (O_functor g) (O_rec_beta (fun x : A => to O B (f x)) a) @ O_rec_beta (fun x : B => to O C (g x)) (f a) = (O_indpaths (O_functor (fun x : A => g (f x))) (fun x : O A => O_functor g (O_functor f x)) (fun x : A => O_rec_beta (fun x0 : A => to O C (g (f x0))) x @ ((O_rec_beta (fun x0 : B => to O C (g x0)) (f x))^ @ ap (O_functor g) (O_rec_beta (fun x0 : A => to O B (f x0)) x)^)) (to O A a))^ @ O_rec_beta (fun x : A => to O C (g (f x))) a
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
a: A

ap (O_functor g) (O_rec_beta (fun x : A => to O B (f x)) a) @ O_rec_beta (fun x : B => to O C (g x)) (f a) = (O_rec_beta (fun x : A => to O C (g (f x))) a @ ((O_rec_beta (fun x : B => to O C (g x)) (f a))^ @ ap (O_functor g) (O_rec_beta (fun x : A => to O B (f x)) a)^))^ @ O_rec_beta (fun x : A => to O C (g (f x))) a
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
a: A

ap (O_functor g) (O_rec_beta (fun x : A => to O B (f x)) a) @ O_rec_beta (fun x : B => to O C (g x)) (f a) = ap (O_functor g) (O_rec_beta (fun x : A => to O B (f x)) a) @ (O_rec_beta (fun x : B => to O C (g x)) (f a) @ ((O_rec_beta (fun x : A => to O C (g (f x))) a)^ @ O_rec_beta (fun x : A => to O C (g (f x))) a))
rewrite concat_Vp, concat_p1; reflexivity. Qed. (** The pseudofunctoriality axiom *)
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: O A

O_functor_compose f (h o g) a @ O_functor_compose g h (O_functor f a) = O_functor_compose (g o f) h a @ ap (O_functor h) (O_functor_compose f g a)
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: O A

O_functor_compose f (h o g) a @ O_functor_compose g h (O_functor f a) = O_functor_compose (g o f) h a @ ap (O_functor h) (O_functor_compose f g a)
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D

(fun x : O_reflector O A => O_functor_compose f (fun x0 : B => h (g x0)) x @ O_functor_compose g h (O_functor f x)) oD to O A == (fun x : O_reflector O A => O_functor_compose (fun x0 : A => g (f x0)) h x @ ap (O_functor h) (O_functor_compose f g x)) oD to O A
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

O_indpaths (O_functor (fun x : A => h (g (f x)))) (fun x : O A => O_functor (fun x0 : B => h (g x0)) (O_functor f x)) (fun x : A => to_O_natural (fun x0 : A => h (g (f x0))) x @ ((to_O_natural (fun x0 : B => h (g x0)) (f x))^ @ ap (O_functor (fun x0 : B => h (g x0))) (to_O_natural f x)^)) (to O A a) @ O_indpaths (O_functor (fun x : B => h (g x))) (fun x : O B => O_functor h (O_functor g x)) (fun x : B => to_O_natural (fun x0 : B => h (g x0)) x @ ((to_O_natural h (g x))^ @ ap (O_functor h) (to_O_natural g x)^)) (O_functor f (to O A a)) = O_indpaths (O_functor (fun x : A => h (g (f x)))) (fun x : O A => O_functor h (O_functor (fun x0 : A => g (f x0)) x)) (fun x : A => to_O_natural (fun x0 : A => h (g (f x0))) x @ ((to_O_natural h (g (f x)))^ @ ap (O_functor h) (to_O_natural (fun x0 : A => g (f x0)) x)^)) (to O A a) @ ap (O_functor h) (O_indpaths (O_functor (fun x : A => g (f x))) (fun x : O A => O_functor g (O_functor f x)) (fun x : A => to_O_natural (fun x0 : A => g (f x0)) x @ ((to_O_natural g (f x))^ @ ap (O_functor g) (to_O_natural f x)^)) (to O A a))
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

O_indpaths (O_functor (fun x : A => h (g (f x)))) (fun x : O A => O_functor (fun x0 : B => h (g x0)) (O_functor f x)) (fun x : A => to_O_natural (fun x0 : A => h (g (f x0))) x @' ((to_O_natural (fun x0 : B => h (g x0)) (f x))^ @' ap (O_functor (fun x0 : B => h (g x0))) (to_O_natural f x)^)) (to O A a) @' O_indpaths (O_functor (fun x : B => h (g x))) (fun x : O B => O_functor h (O_functor g x)) (fun x : B => to_O_natural (fun x0 : B => h (g x0)) x @' ((to_O_natural h (g x))^ @' ap (O_functor h) (to_O_natural g x)^)) (O_functor f (to O A a)) = O_indpaths (O_functor (fun x : A => h (g (f x)))) (fun x : O A => O_functor h (O_functor (fun x0 : A => g (f x0)) x)) (fun x : A => to_O_natural (fun x0 : A => h (g (f x0))) x @' ((to_O_natural h (g (f x)))^ @' ap (O_functor h) (to_O_natural (fun x0 : A => g (f x0)) x)^)) (to O A a) @' ap (O_functor h) (O_indpaths (O_functor (fun x : A => g (f x))) (fun x : O A => O_functor g (O_functor f x)) (fun x : A => to_O_natural (fun x0 : A => g (f x0)) x @' ((to_O_natural g (f x))^ @' ap (O_functor g) (to_O_natural f x)^)) (to O A a))
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural (fun x : B => h (g x)) (f a))^ @' (ap (O_functor (fun x : B => h (g x))) (to_O_natural f a))^ @' O_indpaths (O_functor (fun x : B => h (g x))) (fun x : O B => O_functor h (O_functor g x)) (fun x : B => to_O_natural (fun x0 : B => h (g x0)) x @' ((to_O_natural h (g x))^ @' ap (O_functor h) (to_O_natural g x)^)) (O_functor f (to O A a)) = to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural h (g (f a)))^ @' (ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a))^ @' ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a) @' (ap (O_functor h) (to_O_natural g (f a)))^ @' (ap (O_functor h) (ap (O_functor g) (to_O_natural f a)))^
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural (fun x : B => h (g x)) (f a))^ @' (ap (O_functor (fun x : B => h (g x))) (to_O_natural f a))^ @' transport (fun x : O_reflector O B => O_functor (fun x0 : B => h (g x0)) x = O_functor h (O_functor g x)) (to_O_natural f a)^ (O_indpaths (O_functor (fun x : B => h (g x))) (fun x : O B => O_functor h (O_functor g x)) (fun x : B => to_O_natural (fun x0 : B => h (g x0)) x @' ((to_O_natural h (g x))^ @' ap (O_functor h) (to_O_natural g x)^)) (to O B (f a))) = to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural h (g (f a)))^ @' (ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a))^ @' ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a) @' (ap (O_functor h) (to_O_natural g (f a)))^ @' (ap (O_functor h) (ap (O_functor g) (to_O_natural f a)))^
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural (fun x : B => h (g x)) (f a))^ @' (ap (O_functor (fun x : B => h (g x))) (to_O_natural f a))^ @' transport (fun x : O_reflector O B => O_functor (fun x0 : B => h (g x0)) x = O_functor h (O_functor g x)) (to_O_natural f a)^ (to_O_natural (fun x : B => h (g x)) (f a) @' ((to_O_natural h (g (f a)))^ @' ap (O_functor h) (to_O_natural g (f a))^)) = to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural h (g (f a)))^ @' (ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a))^ @' ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a) @' (ap (O_functor h) (to_O_natural g (f a)))^ @' (ap (O_functor h) (ap (O_functor g) (to_O_natural f a)))^
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural (fun x : B => h (g x)) (f a))^ @' (ap (O_functor (fun x : B => h (g x))) (to_O_natural f a))^ @' (ap (O_functor (fun x : B => h (g x))) (to_O_natural f a)^)^ @' to_O_natural (fun x : B => h (g x)) (f a) @' (to_O_natural h (g (f a)))^ @' ap (O_functor h) (to_O_natural g (f a))^ @' ap (fun x : O B => O_functor h (O_functor g x)) (to_O_natural f a)^ = to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural h (g (f a)))^ @' (ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a))^ @' ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a) @' (ap (O_functor h) (to_O_natural g (f a)))^ @' (ap (O_functor h) (ap (O_functor g) (to_O_natural f a)))^
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural (fun x : B => h (g x)) (f a))^ @' (ap (O_functor (fun x : B => h (g x))) (to_O_natural f a))^ @' ap (O_functor (fun x : B => h (g x))) (to_O_natural f a) @' to_O_natural (fun x : B => h (g x)) (f a) @' (to_O_natural h (g (f a)))^ @' (ap (O_functor h) (to_O_natural g (f a)))^ @' (ap (fun x : O B => O_functor h (O_functor g x)) (to_O_natural f a))^ = to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural h (g (f a)))^ @' (ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a))^ @' ap (O_functor h) (to_O_natural (fun x : A => g (f x)) a) @' (ap (O_functor h) (to_O_natural g (f a)))^ @' (ap (O_functor h) (ap (O_functor g) (to_O_natural f a)))^
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural h (g (f a)))^ @' (ap (O_functor h) (to_O_natural g (f a)))^ @' (ap (fun x : O B => O_functor h (O_functor g x)) (to_O_natural f a))^ = to_O_natural (fun x : A => h (g (f x))) a @' (to_O_natural h (g (f a)))^ @' (ap (O_functor h) (to_O_natural g (f a)))^ @' (ap (O_functor h) (ap (O_functor g) (to_O_natural f a)))^
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

(ap (fun x : O B => O_functor h (O_functor g x)) (to_O_natural f a))^ = (ap (O_functor h) (ap (O_functor g) (to_O_natural f a)))^
O: ReflectiveSubuniverse
A, B, C, D: Type
f: A -> B
g: B -> C
h: C -> D
a: A

ap (fun x : O B => O_functor h (O_functor g x)) (to_O_natural f a) = ap (O_functor h) (ap (O_functor g) (to_O_natural f a))
apply ap_compose. Close Scope long_path_scope. Qed. (** Preservation of equivalences *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

IsEquiv (O_functor f)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

IsEquiv (O_functor f)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

(fun x : O B => O_functor f (O_functor f^-1 x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
(fun x : O A => O_functor f^-1 (O_functor f x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

(fun x : O B => O_functor f (O_functor f^-1 x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O B

O_functor f (O_functor f^-1 x) = x
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O B

O_functor (fun x : B => f (f^-1 x)) x = x
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O B

(fun x : B => f (f^-1 x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O B
O_functor idmap x = x
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O B

(fun x : B => f (f^-1 x)) == idmap
intros y; apply eisretr.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O B

O_functor idmap x = x
apply O_functor_idmap.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

(fun x : O A => O_functor f^-1 (O_functor f x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O A

O_functor f^-1 (O_functor f x) = x
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O A

O_functor (fun x : A => f^-1 (f x)) x = x
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O A

(fun x : A => f^-1 (f x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O A
O_functor idmap x = x
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O A

(fun x : A => f^-1 (f x)) == idmap
intros y; apply eissect.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f
x: O A

O_functor idmap x = x
apply O_functor_idmap. Defined. Definition equiv_O_functor {A B : Type} (f : A <~> B) : O A <~> O B := Build_Equiv _ _ (O_functor f) _. (** This is sometimes useful to have a separate name for, to facilitate rewriting along it. *) Definition to_O_equiv_natural {A B} (f : A <~> B) : (equiv_O_functor f) o (to O A) == (to O B) o f := to_O_natural f. (** This corresponds to [ap O] on the universe. *)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A <~> B

ap O (path_universe_uncurried f) = path_universe_uncurried (equiv_O_functor f)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A <~> B

ap O (path_universe_uncurried f) = path_universe_uncurried (equiv_O_functor f)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type

forall f : A <~> B, ap O (path_universe_uncurried f) = path_universe_uncurried (equiv_O_functor f)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
p: A = B

ap O (path_universe_uncurried (equiv_path A B p)) = path_universe_uncurried (equiv_O_functor (equiv_path A B p))
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
p: A = B

ap O p = path_universe_uncurried (equiv_O_functor (equiv_path A B p))
O: ReflectiveSubuniverse
H: Univalence
A: Type

1 = path_universe_uncurried (equiv_O_functor (equiv_path A A 1))
O: ReflectiveSubuniverse
H: Univalence
A: Type

equiv_path (O A) (O A) 1 = equiv_O_functor (equiv_path A A 1)
O: ReflectiveSubuniverse
H: Univalence
A: Type
x: A

equiv_path (O A) (O A) 1 (to O A x) = equiv_O_functor (equiv_path A A 1) (to O A x)
symmetry; apply to_O_natural. Defined. Definition ap_O_path_universe `{Univalence} {A B : Type} (f : A -> B) `{IsEquiv _ _ f} : ap O (path_universe f) = path_universe (O_functor f) := ap_O_path_universe' (Build_Equiv _ _ f _). (** Postcomposition respects [O_rec] *)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C

g o O_rec f == O_rec (g o f)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C

g o O_rec f == O_rec (g o f)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C
x: A

g (O_rec f (to O A x)) = O_rec (fun x : A => g (f x)) (to O A x)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C
x: A

g (O_rec f (to O A x)) = g (f x)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C
x: A
g (f x) = O_rec (fun x : A => g (f x)) (to O A x)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C
x: A

g (O_rec f (to O A x)) = g (f x)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C
x: A

O_rec f (to O A x) = f x
apply O_rec_beta.
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C
x: A

g (f x) = O_rec (fun x : A => g (f x)) (to O A x)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O B
C_inO: In O C
f: A -> B
g: B -> C
x: A

O_rec (fun x : A => g (f x)) (to O A x) = g (f x)
exact (O_rec_beta (g o f) x). Defined. (** In particular, we have: *) Definition O_rec_postcompose_to_O {A B : Type} (f : A -> B) `{In O B} : to O B o O_rec f == O_functor f := O_rec_postcompose f (to O B). End Functor. Section Replete. (** An equivalent formulation of repleteness is that a type lies in the subuniverse as soon as its unit map is an equivalence. *) Definition inO_isequiv_to_O (T:Type) : IsEquiv (to O T) -> In O T := fun _ => inO_equiv_inO (O T) (to O T)^-1. (** We don't make this an ordinary instance, but we allow it to solve [In O] constraints if we already have [IsEquiv] as a hypothesis. *) #[local] Hint Immediate inO_isequiv_to_O : typeclass_instances.
O: ReflectiveSubuniverse
T: Type

In O T <-> IsEquiv (to O T)
O: ReflectiveSubuniverse
T: Type

In O T <-> IsEquiv (to O T)
split; exact _. Defined. (** Thus, [T] is in a subuniverse as soon as [to O T] admits a retraction. *)
O: ReflectiveSubuniverse
T: Type
mu: O T -> T

mu o to O T == idmap -> In O T
O: ReflectiveSubuniverse
T: Type
mu: O T -> T

mu o to O T == idmap -> In O T
O: ReflectiveSubuniverse
T: Type
mu: O T -> T
H: mu o to O T == idmap

In O T
O: ReflectiveSubuniverse
T: Type
mu: O T -> T
H: mu o to O T == idmap

IsEquiv (to O T)
O: ReflectiveSubuniverse
T: Type
mu: O T -> T
H: mu o to O T == idmap

(fun x : O T => to O T (mu x)) == idmap
O: ReflectiveSubuniverse
T: Type
mu: O T -> T
H: mu o to O T == idmap
(fun x : T => mu (to O T x)) == idmap
O: ReflectiveSubuniverse
T: Type
mu: O T -> T
H: mu o to O T == idmap

(fun x : O T => to O T (mu x)) == idmap
O: ReflectiveSubuniverse
T: Type
mu: O T -> T
H: mu o to O T == idmap

(fun x : T => to O T (mu (to O T x))) == (fun x : T => to O T x)
intros x; exact (ap (to O T) (H x)).
O: ReflectiveSubuniverse
T: Type
mu: O T -> T
H: mu o to O T == idmap

(fun x : T => mu (to O T x)) == idmap
exact H. Defined. (** It follows that reflective subuniverses are closed under retracts. *)
O: ReflectiveSubuniverse
A, B: Type
H: In O B
s: A -> B
r: B -> A
K: r o s == idmap

In O A
O: ReflectiveSubuniverse
A, B: Type
H: In O B
s: A -> B
r: B -> A
K: r o s == idmap

In O A
O: ReflectiveSubuniverse
A, B: Type
H: In O B
s: A -> B
r: B -> A
K: r o s == idmap

(fun x : A => r ((to O B)^-1 (O_functor s (to O A x)))) == idmap
O: ReflectiveSubuniverse
A, B: Type
H: In O B
s: A -> B
r: B -> A
K: r o s == idmap
a: A

r ((to O B)^-1 (O_functor s (to O A a))) = a
O: ReflectiveSubuniverse
A, B: Type
H: In O B
s: A -> B
r: B -> A
K: r o s == idmap
a: A

r ((to O B)^-1 (to O B (s a))) = a
O: ReflectiveSubuniverse
A, B: Type
H: In O B
s: A -> B
r: B -> A
K: r o s == idmap
a: A

r (s a) = a
apply K. Defined. End Replete. (** The maps that are inverted by the reflector. Note that this notation is not (yet) global (because notations in a section cannot be made global); it only exists in this section. After the section is over, we will redefine it globally. *) Local Notation O_inverts f := (IsEquiv (O_functor f)). Section OInverts.
O: ReflectiveSubuniverse
A: Type

O_inverts (to O A)
O: ReflectiveSubuniverse
A: Type

O_inverts (to O A)
O: ReflectiveSubuniverse
A: Type

to O (O_reflector O A) == O_functor (to O A)
intros x; symmetry; apply O_functor_wellpointed. Defined. (** A map between modal types that is inverted by [O] is already an equivalence. This can't be an [Instance], probably because it causes an infinite regress applying more and more [O_functor]. *)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
O_inverts0: O_inverts f

IsEquiv f
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
O_inverts0: O_inverts f

IsEquiv f
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
O_inverts0: O_inverts f

(fun x : A => O_functor f (to O A x)) == (fun x : A => to O B (f x))
apply to_O_natural. Defined. (** Strangely, even this seems to cause infinite loops *) (** [Hint Immediate isequiv_O_inverts : typeclass_instances.] *) Definition equiv_O_inverts {A B : Type} `{In O A} `{In O B} (f : A -> B) `{O_inverts f} : A <~> B := Build_Equiv _ _ f (isequiv_O_inverts f).
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
O_inverts0: O_inverts f

IsEquiv (O_rec f)
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
O_inverts0: O_inverts f

IsEquiv (O_rec f)
(* Not sure why we need [C:=O B] on the next line to get Coq to use two typeclass instances. *)
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
O_inverts0: O_inverts f

IsEquiv (fun x : O_reflector O A => to O B (O_rec f x))
rapply (isequiv_homotopic (O_functor f) (fun x => (O_rec_postcompose_to_O f x)^)). Defined. Definition equiv_O_rec_O_inverts {A B : Type} `{In O B} (f : A -> B) `{O_inverts f} : O A <~> B := Build_Equiv _ _ _ (isequiv_O_rec_O_inverts f). Definition isequiv_to_O_O_inverts {A B : Type} `{In O A} (f : A -> B) `{O_inverts f} : IsEquiv (to O B o f) := isequiv_homotopic (O_functor f o to O A) (to_O_natural f). Definition equiv_to_O_O_inverts {A B : Type} `{In O A} (f : A -> B) `{O_inverts f} : A <~> O B := Build_Equiv _ _ _ (isequiv_to_O_O_inverts f). (** If [f] is inverted by [O], then mapping out of it into any modal type is an equivalence. First we prove a version not requiring funext. For use in [O_inverts_O_leq] below, we allow the types [A], [B], and [Z] to perhaps live in smaller universes than the one [i] on which our subuniverse lives. This the first half of Lemma 1.23 of RSS. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

ooExtendableAlong f (fun _ : B => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

ooExtendableAlong f (fun _ : B => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

ooExtendableAlong (to O B) (fun _ : O_reflector O B => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z
ooExtendableAlong (fun x : A => to O B (f x)) (fun _ : O_reflector O B => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

ooExtendableAlong (fun x : A => to O B (f x)) (fun _ : O_reflector O B => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

(fun x : A => O_functor f (to O A x)) == (fun x : A => to O B (f x))
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z
ooExtendableAlong (fun x : A => O_functor f (to O A x)) (fun _ : O_reflector O B => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

ooExtendableAlong (fun x : A => O_functor f (to O A x)) (fun _ : O_reflector O B => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

ooExtendableAlong (O_functor f) (fun _ : O_reflector O B => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z
ooExtendableAlong (to O A) (fun _ : O_reflector O A => Z)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

ooExtendableAlong (O_functor f) (fun _ : O_reflector O B => Z)
srapply ooextendable_equiv.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H: In O Z

ooExtendableAlong (to O A) (fun _ : O_reflector O A => Z)
exact (extendable_to_O'@{i a z} O A). Defined. (** And now the funext version *)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H0: In O Z

IsEquiv (fun g : B -> Z => g o f)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H0: In O Z

IsEquiv (fun g : B -> Z => g o f)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
O_inverts0: O_inverts f
Z: Type
H0: In O Z

ExtendableAlong 2 f (fun _ : B => Z)
exact (ooextendable_O_inverts f Z 2). Defined. (** Conversely, if a map is inverted by the representable functor [? -> Z] for all [O]-modal types [Z], then it is inverted by [O]. As before, first we prove a version that doesn't require funext. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)

O_inverts f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)

O_inverts f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)

O B -> O A
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
O_functor f o ?g == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
?g o O_functor f == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)

O B -> O A
exact (O_rec (fst (e (O A) _) (to O A)).1).
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)

O_functor f o O_rec (fst (e (O A) (O_inO A)) (to O A)).1 == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)

(fun x : O B => O_functor f (O_rec (fst (e (O A) (O_inO A)) (to O A)).1 x)) o to O B == idmap o to O B
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
b: B

O_functor f (O_rec (fst (e (O A) (O_inO A)) (to O A)).1 (to O B b)) = to O B b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
b: B

O_functor f ((fst (e (O A) (O_inO A)) (to O A)).1 b) = to O B b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
b: B
e1: forall (h k : forall b : B, (fun _ : B => O B) b) (g : forall a : A, (fun b : B => h b = k b) (f a)), ExtensionAlong f (fun b : B => h b = k b) g

O_functor f ((fst (e (O A) (O_inO A)) (to O A)).1 b) = to O B b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
b: B
e1: forall (h k : B -> O B) (g : forall a : A, h (f a) = k (f a)), ExtensionAlong f (fun b : B => h b = k b) g

O_functor f ((fst (e (O A) (O_inO A)) (to O A)).1 b) = to O B b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
b: B
e1: forall (h k : B -> O B) (g : forall a : A, h (f a) = k (f a)), ExtensionAlong f (fun b : B => h b = k b) g

forall a : A, O_functor f ((fst (e (O A) (O_inO A)) (to O A)).1 (f a)) = to O B (f a)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
b: B
e1: forall (h k : B -> O B) (g : forall a : A, h (f a) = k (f a)), ExtensionAlong f (fun b : B => h b = k b) g
a: A

O_functor f ((fst (e (O A) (O_inO A)) (to O A)).1 (f a)) = to O B (f a)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
b: B
e1: forall (h k : B -> O B) (g : forall a : A, h (f a) = k (f a)), ExtensionAlong f (fun b : B => h b = k b) g
a: A

O_functor f (to O A a) = to O B (f a)
apply to_O_natural.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)

O_rec (fst (e (O A) (O_inO A)) (to O A)).1 o O_functor f == idmap
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)

(fun x : O A => O_rec (fst (e (O A) (O_inO A)) (to O A)).1 (O_functor f x)) o to O A == idmap o to O A
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
a: A

O_rec (fst (e (O A) (O_inO A)) (to O A)).1 (O_functor f (to O A a)) = to O A a
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
a: A

(fst (e (O A) (O_inO A)) (to O A)).1 (f a) = to O A a
exact ((fst (e (O A) (O_inO A)) (to O A)).2 a). Defined. #[local] Hint Extern 4 => progress (cbv beta iota) : typeclass_instances. (** And the version with funext. Use it with universe parameters [i j k l lplus l l l l]. *)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> IsEquiv (fun g : B -> Z => g o f)

O_inverts f
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> IsEquiv (fun g : B -> Z => g o f)

O_inverts f
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> IsEquiv (fun g : B -> Z => g o f)

forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
e: forall Z : Type, In O Z -> IsEquiv (fun g : B -> Z => g o f)
Z: Type
X: In O Z

ExtendableAlong 2 f (fun _ : B => Z)
rapply ((equiv_extendable_isequiv 0 _ _)^-1%equiv). Defined. (** This property also characterizes the types in the subuniverse, which is the other half of Lemma 1.23. *)
O: ReflectiveSubuniverse
Z: Type
E: forall (A B : Type) (f : A -> B), O_inverts f -> ooExtendableAlong f (fun _ : B => Z)

In O Z
O: ReflectiveSubuniverse
Z: Type
E: forall (A B : Type) (f : A -> B), O_inverts f -> ooExtendableAlong f (fun _ : B => Z)

In O Z
O: ReflectiveSubuniverse
Z: Type
E: forall (A B : Type) (f : A -> B), O_inverts f -> ooExtendableAlong f (fun _ : B => Z)
EZ:= fst (E Z (O Z) (to O Z) (O_inverts_O_unit Z) 1) idmap: ExtensionAlong (to O Z) (fun _ : O Z => Z) idmap

In O Z
exact (inO_to_O_retract _ EZ.1 EZ.2). Defined. (** A version with the equivalence form of the extension condition. *)
O: ReflectiveSubuniverse
Z: Type
Yo: forall (A B : Type) (f : A -> B), O_inverts f -> IsEquiv (fun g : B -> Z => g o f)

In O Z
O: ReflectiveSubuniverse
Z: Type
Yo: forall (A B : Type) (f : A -> B), O_inverts f -> IsEquiv (fun g : B -> Z => g o f)

In O Z
O: ReflectiveSubuniverse
Z: Type
Yo: forall (A B : Type) (f : A -> B), O_inverts f -> IsEquiv (fun g : B -> Z => g o f)
EZ:= extension_isequiv_precompose (to O Z) (fun _ : O Z => Z) (Yo Z (O Z) (to O Z) (O_inverts_O_unit Z)) idmap: ExtensionAlong (to O Z) (fun _ : O Z => Z) idmap

In O Z
exact (inO_to_O_retract _ EZ.1 EZ.2). Defined.
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B

(to O B)^-1 o O_functor f == f o (to O A)^-1
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B

(to O B)^-1 o O_functor f == f o (to O A)^-1
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
x: A

(to O B)^-1 (O_functor f (to O A x)) = f ((to O A)^-1 (to O A x))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
x: A

O_functor f (to O A x) = to O B (f ((to O A)^-1 (to O A x)))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
x: A

to O B (f x) = to O B (f ((to O A)^-1 (to O A x)))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
x: A

x = (to O A)^-1 (to O A x)
symmetry; apply eissect. Defined. (** Two maps between modal types that become equal after applying [O_functor] are already equal. *)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g

f == g
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g

f == g
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

f x = g x
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

f ((to O A)^-1 (to O A x)) = g x
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

f ((to O A)^-1 (to O A x)) = g ((to O A)^-1 (to O A x))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

f ((to O A)^-1 (to O A x)) = (to O B)^-1 (O_functor f (to O A x))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A
(to O B)^-1 (O_functor f (to O A x)) = g ((to O A)^-1 (to O A x))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

f ((to O A)^-1 (to O A x)) = (to O B)^-1 (O_functor f (to O A x))
symmetry; apply to_O_inv_natural.
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

(to O B)^-1 (O_functor f (to O A x)) = g ((to O A)^-1 (to O A x))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

(to O B)^-1 (O_functor f (to O A x)) = (to O B)^-1 (O_functor g (to O A x))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A
(to O B)^-1 (O_functor g (to O A x)) = g ((to O A)^-1 (to O A x))
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

(to O B)^-1 (O_functor f (to O A x)) = (to O B)^-1 (O_functor g (to O A x))
apply ap, e.
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f, g: A -> B
e: O_functor f == O_functor g
x: A

(to O B)^-1 (O_functor g (to O A x)) = g ((to O A)^-1 (to O A x))
apply to_O_inv_natural. Defined. (** Any map to a type in the subuniverse that is inverted by [O] must be equivalent to [to O]. More precisely, the type of such maps is contractible. *) Definition typeof_to_O (A : Type) := { OA : Type & { Ou : A -> OA & ((In O OA) * (O_inverts Ou)) }}.
O: ReflectiveSubuniverse
H: Univalence
A: Type

Contr (typeof_to_O A)
O: ReflectiveSubuniverse
H: Univalence
A: Type

Contr (typeof_to_O A)
O: ReflectiveSubuniverse
H: Univalence
A: Type

forall y : {OA : Type & {Ou : A -> OA & In O OA * O_inverts Ou}}, (O A; to O A; (O_inO A, O_inverts_O_unit A)) = y
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou

(O A; to O A; (O_inO A, O_inverts_O_unit A)) = (OA; Ou; (fst, snd))
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA

(O A; to O A; (O_inO A, O_inverts_O_unit A)) = (OA; Ou; (fst, snd))
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A

(O A; to O A; (O_inO A, O_inverts_O_unit A)) = (OA; Ou; (fst, snd))
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A

IsEquiv f
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f
(O A; to O A; (O_inO A, O_inverts_O_unit A)) = (OA; Ou; (fst, snd))
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A

IsEquiv f
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A

(fun x : OA => f (g x)) == idmap
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
(fun x : O A => g (f x)) == idmap
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A

(fun x : OA => f (g x)) == idmap
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA

O_functor (fun x : OA => f (g x)) x = O_functor idmap x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA

O_functor (fun x : OA => f (g x)) x = x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA

O_functor f (O_functor g x) = x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA

O_functor f (O_functor (fun x : OA => (O_functor Ou)^-1 (to O OA x)) x) = x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA

O_functor f (O_functor (O_functor Ou)^-1 (O_functor (to O OA) x)) = x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA

O_functor f (O_functor (O_functor Ou)^-1 (to O (O_reflector O OA) x)) = x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA

O_functor f (to O (O A) ((O_functor Ou)^-1 x)) = x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA

to O OA (f ((O_functor Ou)^-1 x)) = x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA
y:= (O_functor Ou)^-1 x: O A

to O OA (f y) = x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA
y:= (O_functor Ou)^-1 x: O A

to O OA (f y) = O_functor Ou y
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: O OA
y:= (O_functor Ou)^-1 x: O A

to O OA (O_rec Ou y) = O_rec (fun x : A => to O OA (Ou x)) y
apply O_rec_postcompose.
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A

(fun x : O A => g (f x)) == idmap
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: A

g (f (to O A x)) = to O A x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: A

g (O_rec Ou (to O A x)) = to O A x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: A

g (Ou x) = to O A x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: A

(O_functor Ou)^-1 (to O OA (Ou x)) = to O A x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
x: A

to O OA (Ou x) = O_functor Ou (to O A x)
symmetry; apply to_O_natural.
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f

(O A; to O A; (O_inO A, O_inverts_O_unit A)) = (OA; Ou; (fst, snd))
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f

O A = OA
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f
transport (fun OA : Type => {Ou : A -> OA & In O OA * O_inverts Ou}) ?Goal (to O A; (O_inO A, O_inverts_O_unit A)) = (Ou; (fst, snd))
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f

O A = OA
exact (path_universe f).
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f

transport (fun OA : Type => {Ou : A -> OA & In O OA * O_inverts Ou}) (path_universe f) (to O A; (O_inO A, O_inverts_O_unit A)) = (Ou; (fst, snd))
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f

(transport (fun x : Type => A -> x) (path_universe f) (to O A; (O_inO A, O_inverts_O_unit A)).1; transportD (fun x : Type => A -> x) (fun (x : Type) (y : A -> x) => In O x * O_inverts y) (path_universe f) (to O A; (O_inO A, O_inverts_O_unit A)).1 (to O A; (O_inO A, O_inverts_O_unit A)).2) = (Ou; (fst, snd))
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f

transport (fun x : Type => A -> x) (path_universe f) (to O A) = Ou
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f
x: A

transport (fun x : Type => A -> x) (path_universe f) (to O A) x = Ou x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f
x: A

transport idmap (path_universe f) (to O A x) = Ou x
O: ReflectiveSubuniverse
H: Univalence
A, OA: Type
Ou: A -> OA
fst: In O OA
snd: O_inverts Ou
f:= O_rec Ou : O A -> OA: O A -> OA
g:= (O_functor Ou)^-1 o to O OA : OA -> O A: OA -> O A
X: IsEquiv f
x: A

f (to O A x) = Ou x
unfold f; apply O_rec_beta. Qed. End OInverts. Section Types. (** ** The [Unit] type *)
O: ReflectiveSubuniverse

In O Unit
O: ReflectiveSubuniverse

In O Unit
O: ReflectiveSubuniverse

unit_name tt == idmap
exact (@contr@{Set} Unit _). Defined. (** It follows that any contractible type is in [O]. *)
O: ReflectiveSubuniverse
A: Type
Contr0: Contr A

In O A
O: ReflectiveSubuniverse
A: Type
Contr0: Contr A

In O A
exact (inO_equiv_inO@{Set _ _} Unit equiv_contr_unit^-1). Defined. (** And that the reflection of a contractible type is still contractible. *)
O: ReflectiveSubuniverse
A: Type
Contr0: Contr A

Contr (O A)
O: ReflectiveSubuniverse
A: Type
Contr0: Contr A

Contr (O A)
exact (contr_equiv A (to O A)). Defined. (** ** Dependent product and arrows *) (** Theorem 7.7.2 *)
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type

(forall x : A, In O (B x)) -> In O (forall x : A, B x)
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type

(forall x : A, In O (B x)) -> In O (forall x : A, B x)
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type
H: forall x : A, In O (B x)

In O (forall x : A, B x)
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type
H: forall x : A, In O (B x)
ev:= fun (x : A) (f : forall x0 : A, B x0) => f x: forall x : A, (forall x0 : A, B x0) -> B x

In O (forall x : A, B x)
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type
H: forall x : A, In O (B x)
ev:= fun (x : A) (f : forall x0 : A, B x0) => f x: forall x : A, (forall x0 : A, B x0) -> B x
zz:= fun x : A => O_rec (ev x): forall x : A, O_reflector O (forall x0 : A, B x0) -> B x

In O (forall x : A, B x)
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type
H: forall x : A, In O (B x)
ev:= fun (x : A) (f : forall x0 : A, B x0) => f x: forall x : A, (forall x0 : A, B x0) -> B x
zz:= fun x : A => O_rec (ev x): forall x : A, O_reflector O (forall x0 : A, B x0) -> B x

(fun (x : forall x : A, B x) (x0 : A) => zz x0 (to O (forall x1 : A, B x1) x)) == idmap
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type
H: forall x : A, In O (B x)
ev:= fun (x : A) (f : forall x0 : A, B x0) => f x: forall x : A, (forall x0 : A, B x0) -> B x
zz:= fun x : A => O_rec (ev x): forall x : A, O_reflector O (forall x0 : A, B x0) -> B x
phi: forall x : A, B x

(fun x : A => zz x (to O (forall x0 : A, B x0) phi)) = phi
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type
H: forall x : A, In O (B x)
phi: forall x : A, B x

(fun x : A => O_rec (fun f : forall x0 : A, B x0 => f x) (to O (forall x0 : A, B x0) phi)) = phi
O: ReflectiveSubuniverse
fs: Funext
A: Type
B: A -> Type
H: forall x : A, In O (B x)
phi: forall x : A, B x
x: A

O_rec (fun f : forall x : A, B x => f x) (to O (forall x : A, B x) phi) = phi x
exact (O_rec_beta (fun f : forall x0, (B x0) => f x) phi). Defined.
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
H: In O B

In O (A -> B)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
H: In O B

In O (A -> B)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
H: In O B

A -> In O B
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
H: In O B
a: A

In O B
exact _. Defined. (** ** Product *)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B

In O (A * B)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B

In O (A * B)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B

(fun x : A * B => (O_rec fst (to O (A * B) x), O_rec snd (to O (A * B) x))) == idmap
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
a: A
b: B

O_rec fst (to O (A * B) (a, b)) = a
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
a: A
b: B
O_rec snd (to O (A * B) (a, b)) = b
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
a: A
b: B

O_rec fst (to O (A * B) (a, b)) = a
exact (O_rec_beta fst (a,b)).
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
a: A
b: B

O_rec snd (to O (A * B) (a, b)) = b
exact (O_rec_beta snd (a,b)). Defined. (** Two ways to define a map [O(A * B) -> X * Y] agree. *)
O: ReflectiveSubuniverse
A, B, X, Y: Type
H: In O X
H0: In O Y
f: A -> X
g: B -> Y

O_rec (functor_prod f g) == prod_coind (O_rec (f o fst : A * B -> X)) (O_rec (g o snd : A * B -> Y))
O: ReflectiveSubuniverse
A, B, X, Y: Type
H: In O X
H0: In O Y
f: A -> X
g: B -> Y

O_rec (functor_prod f g) == prod_coind (O_rec (f o fst : A * B -> X)) (O_rec (g o snd : A * B -> Y))
O: ReflectiveSubuniverse
A, B, X, Y: Type
H: In O X
H0: In O Y
f: A -> X
g: B -> Y
ab: A * B

O_rec (functor_prod f g) (to O (A * B) ab) = prod_coind (O_rec (fun x : A * B => f (fst x))) (O_rec (fun x : A * B => g (snd x))) (to O (A * B) ab)
O: ReflectiveSubuniverse
A, B, X, Y: Type
H: In O X
H0: In O Y
f: A -> X
g: B -> Y
ab: A * B

O_rec (fun z : A * B => (f (fst z), g (snd z))) (to O (A * B) ab) = (O_rec (fun x : A * B => f (fst x)) (to O (A * B) ab), O_rec (fun x : A * B => g (snd x)) (to O (A * B) ab))
O: ReflectiveSubuniverse
A, B, X, Y: Type
H: In O X
H0: In O Y
f: A -> X
g: B -> Y
ab: A * B

(f (fst ab), g (snd ab)) = (O_rec (fun x : A * B => f (fst x)) (to O (A * B) ab), O_rec (fun x : A * B => g (snd x)) (to O (A * B) ab))
apply path_prod; cbn; symmetry; nrapply O_rec_beta. Defined. (** We show that [OA*OB] has the same universal property as [O(A*B)] *) (** Here is the map witnessing the universal property. *) Definition O_prod_unit (A B : Type) : A * B -> O A * O B := functor_prod (to O A) (to O B). (** We express the universal property without funext, using extensions. *)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O C

ooExtendableAlong (O_prod_unit A B) (fun _ : O A * O B => C)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O C

ooExtendableAlong (O_prod_unit A B) (fun _ : O A * O B => C)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O C

O B -> ooExtendableAlong (to O A) (fun _ : O A => C)
O: ReflectiveSubuniverse
A, B, C: Type
H: In O C
O A -> ooExtendableAlong (to O B) (fun _ : O B => C)
all:intros; rapply extendable_to_O. Defined. (** Here's the version with funext. *)
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
H: In O C

IsEquiv (fun f : O A * O B -> C => f o O_prod_unit A B)
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
H: In O C

IsEquiv (fun f : O A * O B -> C => f o O_prod_unit A B)
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
H: In O C

ooExtendableAlong (O_prod_unit A B) (fun _ : O A * O B => C)
rapply ooextendable_O_prod_unit. Defined. Definition equiv_O_prod_unit_precompose {fs : Funext} (A B C : Type) `{In O C} : ((O A) * (O B) -> C) <~> (A * B -> C) := Build_Equiv _ _ _ (isequiv_O_prod_unit_precompose A B C). (** The (funext-free) universal property implies that [O_prod_unit] is an [O]-equivalence, hence induces an equivalence between [O (A*B)] and [O A * O B]. *)
O: ReflectiveSubuniverse
A, B: Type

O_inverts (O_prod_unit A B)
O: ReflectiveSubuniverse
A, B: Type

O_inverts (O_prod_unit A B)
O: ReflectiveSubuniverse
A, B: Type

forall Z : Type, In O Z -> ExtendableAlong 2 (O_prod_unit A B) (fun _ : O A * O B => Z)
intros; rapply ooextendable_O_prod_unit. Defined. Definition O_prod_cmp (A B : Type) : O (A * B) -> O A * O B := O_rec (O_prod_unit A B).
O: ReflectiveSubuniverse
A, B: Type

IsEquiv (O_prod_cmp A B)
O: ReflectiveSubuniverse
A, B: Type

IsEquiv (O_prod_cmp A B)
rapply isequiv_O_rec_O_inverts. Defined. Definition equiv_O_prod_cmp (A B : Type) : O (A * B) <~> (O A * O B) := Build_Equiv _ _ (O_prod_cmp A B) _.
O: ReflectiveSubuniverse
X, Y: Type
x0, x1: X
y0, y1: Y

to O (X * Y) (x0, y0) = to O (X * Y) (x1, y1) <~> (to O X x0 = to O X x1) * (to O Y y0 = to O Y y1)
O: ReflectiveSubuniverse
X, Y: Type
x0, x1: X
y0, y1: Y

to O (X * Y) (x0, y0) = to O (X * Y) (x1, y1) <~> (to O X x0 = to O X x1) * (to O Y y0 = to O Y y1)
O: ReflectiveSubuniverse
X, Y: Type
x0, x1: X
y0, y1: Y

equiv_O_prod_cmp X Y (to O (X * Y) (x0, y0)) = equiv_O_prod_cmp X Y (to O (X * Y) (x1, y1)) <~> (to O X x0 = to O X x1) * (to O Y y0 = to O Y y1)
O: ReflectiveSubuniverse
X, Y: Type
x0, x1: X
y0, y1: Y

?Goal0 = ?Goal1 <~> (to O X x0 = to O X x1) * (to O Y y0 = to O Y y1)
O: ReflectiveSubuniverse
X, Y: Type
x0, x1: X
y0, y1: Y
equiv_O_prod_cmp X Y (to O (X * Y) (x0, y0)) = ?Goal0
O: ReflectiveSubuniverse
X, Y: Type
x0, x1: X
y0, y1: Y
equiv_O_prod_cmp X Y (to O (X * Y) (x1, y1)) = ?Goal1
O: ReflectiveSubuniverse
X, Y: Type
x0, x1: X
y0, y1: Y

O_prod_unit X Y (x0, y0) = O_prod_unit X Y (x1, y1) <~> (to O X x0 = to O X x1) * (to O Y y0 = to O Y y1)
exact (equiv_path_prod _ _)^-1%equiv. Defined. Definition O_prod_cmp_coind (A B : Type) : O_prod_cmp A B == prod_coind (O_rec (to O _ o fst : A * B -> O A)) (O_rec (to O _ o snd : A * B -> O B)) := O_rec_functor_prod _ _. (** ** Pullbacks *)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C

In O (Pullback f g)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C

In O (Pullback f g)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C

O (Pullback f g) -> Pullback f g
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
?mu o to O (Pullback f g) == idmap
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C

O (Pullback f g) -> Pullback f g
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
op: O (Pullback f g)

Pullback f g
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
op: O (Pullback f g)

{c : C & f (O_rec pr1 op) = g c}
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
op: O (Pullback f g)

f (O_rec pr1 op) = g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) op)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

f (O_rec pr1 (to O (Pullback f g) (b; c; a))) = g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)))
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

f b = g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)))
refine (a @ ap g (O_rec_beta _ _)^).
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C

(fun op : O (Pullback f g) => (O_rec pr1 op; O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) op; O_indpaths (fun x : O_reflector O (Pullback f g) => f (O_rec pr1 x)) (fun x : O_reflector O (Pullback f g) => g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) x)) ((fun x : Pullback f g => (fun (b : B) (proj2 : {c : C & f b = g c}) => (fun (c : C) (a : f b = g c) => ap f (O_rec_beta pr1 (b; c; a)) @ (a @ ap g (O_rec_beta (fun x0 : ... => (x0.2).1) (b; c; a))^ : f (b; c; a).1 = g (O_rec (fun p : ... => (p.2).1) (to O (Pullback f g) (b; c; a))))) proj2.1 proj2.2) x.1 x.2) : (fun x : Pullback f g => f (O_rec pr1 (to O (Pullback f g) x))) == (fun x : Pullback f g => g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) x)))) op)) o to O (Pullback f g) == idmap
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

(O_rec pr1 (to O (Pullback f g) (b; c; a)); O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)); O_indpaths (fun x : O_reflector O (Pullback f g) => f (O_rec pr1 x)) (fun x : O_reflector O (Pullback f g) => g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) x)) (fun x : Pullback f g => ap f (O_rec_beta pr1 (x.1; (x.2).1; (x.2).2)) @ ((x.2).2 @ ap g (O_rec_beta (fun x0 : Pullback f g => (x0.2).1) (x.1; (x.2).1; (x.2).2))^)) (to O (Pullback f g) (b; c; a))) = (b; c; a)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

O_rec pr1 (to O (Pullback f g) (b; c; a)) = b
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c
transport (fun b : B => {c : C & f b = g c}) ?p (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)); O_indpaths (fun x : O_reflector O (Pullback f g) => f (O_rec pr1 x)) (fun x : O_reflector O (Pullback f g) => g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) x)) (fun x : Pullback f g => ap f (O_rec_beta pr1 (x.1; (x.2).1; (x.2).2)) @ ((x.2).2 @ ap g (O_rec_beta (fun x0 : Pullback f g => (x0.2).1) (x.1; (x.2).1; (x.2).2))^)) (to O (Pullback f g) (b; c; a))) = (c; a)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

O_rec pr1 (to O (Pullback f g) (b; c; a)) = b
apply O_rec_beta.
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

transport (fun b : B => {c : C & f b = g c}) (O_rec_beta pr1 (b; c; a)) (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)); O_indpaths (fun x : O_reflector O (Pullback f g) => f (O_rec pr1 x)) (fun x : O_reflector O (Pullback f g) => g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) x)) (fun x : Pullback f g => ap f (O_rec_beta pr1 (x.1; (x.2).1; (x.2).2)) @ ((x.2).2 @ ap g (O_rec_beta (fun x0 : Pullback f g => (x0.2).1) (x.1; (x.2).1; (x.2).2))^)) (to O (Pullback f g) (b; c; a))) = (c; a)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

(O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)); transport (fun x : B => f x = g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)))) (O_rec_beta pr1 (b; c; a)) (O_indpaths (fun x : O_reflector O (Pullback f g) => f (O_rec pr1 x)) (fun x : O_reflector O (Pullback f g) => g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) x)) (fun x : Pullback f g => ap f (O_rec_beta pr1 (x.1; (x.2).1; (x.2).2)) @ ((x.2).2 @ ap g (O_rec_beta (fun x0 : Pullback f g => (x0.2).1) (x.1; (x.2).1; (x.2).2))^)) (to O (Pullback f g) (b; c; a)))) = (c; a)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)) = c
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c
transport (fun y : C => f b = g y) ?p (transport (fun x : B => f x = g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)))) (O_rec_beta pr1 (b; c; a)) (O_indpaths (fun x : O_reflector O (Pullback f g) => f (O_rec pr1 x)) (fun x : O_reflector O (Pullback f g) => g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) x)) (fun x : Pullback f g => ap f (O_rec_beta pr1 (x.1; (x.2).1; (x.2).2)) @ ((x.2).2 @ ap g (O_rec_beta (fun x0 : Pullback f g => (x0.2).1) (x.1; (x.2).1; (x.2).2))^)) (to O (Pullback f g) (b; c; a)))) = a
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)) = c
apply O_rec_beta.
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: In O A
H0: In O B
H1: In O C
b: B
c: C
a: f b = g c

transport (fun y : C => f b = g y) (O_rec_beta (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (b; c; a)) (transport (fun x : B => f x = g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) (to O (Pullback f g) (b; c; a)))) (O_rec_beta pr1 (b; c; a)) (O_indpaths (fun x : O_reflector O (Pullback f g) => f (O_rec pr1 x)) (fun x : O_reflector O (Pullback f g) => g (O_rec (fun p : {b : B & {c : C & f b = g c}} => (p.2).1) x)) (fun x : Pullback f g => ap f (O_rec_beta pr1 (x.1; (x.2).1; (x.2).2)) @ ((x.2).2 @ ap g (O_rec_beta (fun x0 : Pullback f g => (x0.2).1) (x.1; (x.2).1; (x.2).2))^)) (to O (Pullback f g) (b; c; a)))) = a
abstract ( rewrite transport_paths_Fr; rewrite transport_paths_Fl; rewrite O_indpaths_beta; rewrite concat_V_pp; rewrite ap_V; apply concat_pV_p ). Defined. (** ** Fibers *)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B

In O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B

In O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B

O (hfiber f b) -> hfiber f b
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
?mu o to O (hfiber f b) == idmap
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B

O (hfiber f b) -> hfiber f b
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
x: O (hfiber f b)

A
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
x: O (hfiber f b)
(fun x : A => f x = b) ?proj1
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
x: O (hfiber f b)

A
exact (O_rec pr1 x).
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
x: O (hfiber f b)

(fun x : A => f x = b) (O_rec pr1 x)
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
x: hfiber f b

f (O_rec pr1 (to O (hfiber f b) x)) = b
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
x: hfiber f b

f x.1 = b
exact (x.2).
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B

(fun x : O (hfiber f b) => (O_rec pr1 x; O_indpaths (fun x0 : O_reflector O (hfiber f b) => f (O_rec pr1 x0)) (fun _ : O_reflector O (hfiber f b) => b) ((fun x0 : hfiber f b => ap f (O_rec_beta pr1 x0) @ x0.2 : f (O_rec pr1 (to O (hfiber f b) x0)) = b) : (fun x0 : hfiber f b => f (O_rec pr1 (to O (hfiber f b) x0))) == (fun _ : hfiber f b => b)) x)) o to O (hfiber f b) == idmap
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
a: A
p: f a = b

O_rec pr1 (to O (hfiber f b) (a; p)) = a
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
a: A
p: f a = b
transport (fun x : A => f x = b) ?p (O_indpaths (fun x : O_reflector O (hfiber f b) => f (O_rec pr1 x)) (fun _ : O_reflector O (hfiber f b) => b) (fun x : hfiber f b => ap f (O_rec_beta pr1 x) @ x.2) (to O (hfiber f b) (a; p))) = p
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
a: A
p: f a = b

O_rec pr1 (to O (hfiber f b) (a; p)) = a
exact (O_rec_beta pr1 (a;p)).
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
a: A
p: f a = b

transport (fun x : A => f x = b) (O_rec_beta pr1 (a; p)) (O_indpaths (fun x : O_reflector O (hfiber f b) => f (O_rec pr1 x)) (fun _ : O_reflector O (hfiber f b) => b) (fun x : hfiber f b => ap f (O_rec_beta pr1 x) @ x.2) (to O (hfiber f b) (a; p))) = p
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
a: A
p: f a = b

transport (fun x : A => f x = b) (O_rec_beta pr1 (a; p)) (ap f (O_rec_beta pr1 (a; p)) @ p) = p
O: ReflectiveSubuniverse
A, B: Type
H: In O A
H0: In O B
f: A -> B
b: B
a: A
p: f a = b

(ap f (O_rec_beta pr1 (a; p)))^ @ (ap f (O_rec_beta pr1 (a; p)) @ p) = p
apply concat_V_pp. Defined. Definition inO_unsigma {A : Type} (B : A -> Type) `{In O A} {B_inO : In O {x:A & B x}} (x : A) : In O (B x) := inO_equiv_inO _ (hfiber_fibration x B)^-1. #[local] Hint Immediate inO_unsigma : typeclass_instances. (** The reflector preserving hfibers is a characterization of lex modalities. Here is the comparison map. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B

O (hfiber f b) -> hfiber (O_functor f) (to O B b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B

O (hfiber f b) -> hfiber (O_functor f) (to O B b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B

hfiber f b -> hfiber (O_functor f) (to O B b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B
a: A
p: f a = b

hfiber (O_functor f) (to O B b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B
a: A
p: f a = b

O_functor f (to O A a) = to O B b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B
a: A
p: f a = b

to O B (f a) = to O B b
apply ap, p. Defined.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B

O_functor_hfiber f b o to O (hfiber f b) == functor_hfiber (fun u : A => (to_O_natural f u)^) b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B

O_functor_hfiber f b o to O (hfiber f b) == functor_hfiber (fun u : A => (to_O_natural f u)^) b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B
a: A
p: f a = b

O_rec (fun X : hfiber f b => (to O A X.1; to_O_natural f X.1 @ ap (to O B) X.2)) (to O (hfiber f b) (a; p)) = (to O A a; ((to_O_natural f a)^)^ @ ap (to O B) p)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
b: B
a: A
p: f a = b

(to O A a; to_O_natural f a @ ap (to O B) p) = (to O A a; ((to_O_natural f a)^)^ @ ap (to O B) p)
exact (ap _ (inv_V _ @@ 1))^. Defined. (** [functor_sigma] over [idmap] preserves [O]-equivalences. *)
O: ReflectiveSubuniverse
A: Type
P, Q: A -> Type
g: forall a : A, P a -> Q a
H: forall a : A, O_inverts (g a)

O_inverts (functor_sigma idmap g)
O: ReflectiveSubuniverse
A: Type
P, Q: A -> Type
g: forall a : A, P a -> Q a
H: forall a : A, O_inverts (g a)

O_inverts (functor_sigma idmap g)
O: ReflectiveSubuniverse
A: Type
P, Q: A -> Type
g: forall a : A, P a -> Q a
H: forall a : A, O_inverts (g a)
Z: Type
Z_inO: In O Z

ExtendableAlong 2 (functor_sigma idmap g) (fun _ : {x : _ & Q x} => Z)
O: ReflectiveSubuniverse
A: Type
P, Q: A -> Type
g: forall a : A, P a -> Q a
H: forall a : A, O_inverts (g a)
Z: Type
Z_inO: In O Z
a: A

ooExtendableAlong (g a) (fun _ : Q a => Z)
nrapply ooextendable_O_inverts; exact _. Defined. (** Theorem 7.3.9: The reflector [O] can be discarded inside a reflected sum. This can be obtained from [O_inverts_functor_sigma_id] applied to the family of units [to O (P x)], but unfortunately the definitional behavior of the inverse obtained thereby (which here we take as the "forwards" direction) is poor. So instead we give an explicit proof, but note that the "backwards" direction here is precisely [functor_sigma]. *)
O: ReflectiveSubuniverse
A: Type
P: A -> Type

O {x : A & O (P x)} <~> O {x : A & P x}
(** := (Build_Equiv _ _ _ (O_inverts_functor_sigma_id (fun x => to O (P x))))^-1. *)
O: ReflectiveSubuniverse
A: Type
P: A -> Type

O {x : A & O (P x)} <~> O {x : A & P x}
O: ReflectiveSubuniverse
A: Type
P: A -> Type

O {x : A & O (P x)} -> O {x : A & P x}
O: ReflectiveSubuniverse
A: Type
P: A -> Type
O {x : A & P x} -> O {x : A & O (P x)}
O: ReflectiveSubuniverse
A: Type
P: A -> Type
?f o ?g == idmap
O: ReflectiveSubuniverse
A: Type
P: A -> Type
?g o ?f == idmap
O: ReflectiveSubuniverse
A: Type
P: A -> Type

O {x : A & O (P x)} -> O {x : A & P x}
O: ReflectiveSubuniverse
A: Type
P: A -> Type
a: A

O (P a) -> O {x : A & P x}
O: ReflectiveSubuniverse
A: Type
P: A -> Type
a: A
p: P a

O {x : A & P x}
exact (to O _ (a;p)).
O: ReflectiveSubuniverse
A: Type
P: A -> Type

O {x : A & P x} -> O {x : A & O (P x)}
O: ReflectiveSubuniverse
A: Type
P: A -> Type

{x : A & P x} -> {x : A & O (P x)}
exact (functor_sigma idmap (fun x => to O (P x))).
O: ReflectiveSubuniverse
A: Type
P: A -> Type

O_rec (fun X : {x : A & O (P x)} => (fun (a : A) (op : O (P a)) => O_rec (fun p : P a => to O {x : A & P x} (a; p)) op) X.1 X.2) o O_functor (functor_sigma idmap (fun x : A => to O (P (idmap x)))) == idmap
O: ReflectiveSubuniverse
A: Type
P: A -> Type

(fun x : {x : A & P x} => O_rec (fun X : {x0 : A & O (P x0)} => O_rec (fun p : P X.1 => to O {x0 : A & P x0} (X.1; p)) X.2) (O_rec (fun x0 : {x0 : A & P x0} => to O {x1 : A & O (P x1)} (functor_sigma idmap (fun x1 : A => to O (P x1)) x0)) (to O {x0 : A & P x0} x))) == (fun x : {x : A & P x} => to O {x0 : A & P x0} x)
O: ReflectiveSubuniverse
A: Type
P: A -> Type
a: A
p: P a

O_rec (fun X : {x : A & O (P x)} => O_rec (fun p : P X.1 => to O {x : A & P x} (X.1; p)) X.2) (O_rec (fun x : {x : A & P x} => to O {x0 : A & O (P x0)} (functor_sigma idmap (fun x0 : A => to O (P x0)) x)) (to O {x : A & P x} (a; p))) = to O {x : A & P x} (a; p)
abstract (repeat (rewrite O_rec_beta); reflexivity).
O: ReflectiveSubuniverse
A: Type
P: A -> Type

O_functor (functor_sigma idmap (fun x : A => to O (P (idmap x)))) o O_rec (fun X : {x : A & O (P x)} => (fun (a : A) (op : O (P a)) => O_rec (fun p : P a => to O {x : A & P x} (a; p)) op) X.1 X.2) == idmap
O: ReflectiveSubuniverse
A: Type
P: A -> Type

(fun x : {x : A & O (P x)} => O_rec (fun x0 : {x0 : A & P x0} => to O {x1 : A & O (P x1)} (functor_sigma idmap (fun x1 : A => to O (P x1)) x0)) (O_rec (fun X : {x0 : A & O (P x0)} => O_rec (fun p : P X.1 => to O {x0 : A & P x0} (X.1; p)) X.2) (to O {x0 : A & O (P x0)} x))) == (fun x : {x : A & O (P x)} => to O {x0 : A & O (P x0)} x)
O: ReflectiveSubuniverse
A: Type
P: A -> Type
a: A
p: P a

O_rec (fun x : {x : A & P x} => to O {x0 : A & O (P x0)} (functor_sigma idmap (fun x0 : A => to O (P x0)) x)) (O_rec (fun X : {x : A & O (P x)} => O_rec (fun p : P X.1 => to O {x : A & P x} (X.1; p)) X.2) (to O {x : A & O (P x)} (a; to O (P a) p))) = to O {x : A & O (P x)} (a; to O (P a) p)
abstract (repeat (rewrite O_rec_beta); reflexivity). Defined. (** ** Equivalences *) (** Naively it might seem that we need closure under Sigmas (hence a modality) to deduce closure under [Equiv], but in fact the above closure under fibers is sufficient. This appears as part of the proof of Proposition 2.18 of CORS. For later use, we try to reduce the number of universe parameters (but we don't completely control them all). *)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B

In O (A <~> B)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B

In O (A <~> B)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B

In O {f : A -> B & IsEquiv f}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B

In O {x : _ & BiInv x}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B

(A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= ?Goal: (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
In O {x : _ & BiInv x}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B

(A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
intros [f [g h]]; exact (h o f, f o g).
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun (f : A -> B) (snd : (B -> A) * (B -> A)) => (fun g h : B -> A => (h o f, f o g)) (fst snd) (Overture.snd snd)) (fst X) (snd X): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)

In O {x : _ & BiInv x}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun (f : A -> B) (snd : (B -> A) * (B -> A)) => (fun g h : B -> A => (h o f, f o g)) (fst snd) (Overture.snd snd)) (fst X) (snd X): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type

In O {x : _ & BiInv x}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun (f : A -> B) (snd : (B -> A) * (B -> A)) => (fun g h : B -> A => (h o f, f o g)) (fst snd) (Overture.snd snd)) (fst X) (snd X): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type

U <~> {x : _ & BiInv x}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type

U <~> {f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type

U -> {f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
{f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}} -> U
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
?f o ?g == idmap
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
?g o ?f == idmap
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type

U -> {f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g, h: B -> A
p: c (f, (g, h)) = (idmap, idmap)

{f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g, h: B -> A
p: (fst (c (f, (g, h))) = fst (idmap, idmap)) * (snd (c (f, (g, h))) = snd (idmap, idmap))

{f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g, h: B -> A
p: (fun x : A => h (f x)) = idmap
q: (fun x : B => f (g x)) = idmap

{f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g, h: B -> A
p: (fun x : A => h (f x)) = idmap
q: (fun x : B => f (g x)) = idmap

(fun x : A => h (f x)) == idmap
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g, h: B -> A
p: (fun x : A => h (f x)) = idmap
q: (fun x : B => f (g x)) = idmap
(fun x : B => f (g x)) == idmap
all:apply ap10; assumption.
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type

{f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}} -> U
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap

U
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap

c (f, (h, g)) = (idmap, idmap)
apply path_prod; apply path_arrow; assumption.
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type

(fun X : U => (fun proj1 : (A -> B) * ((B -> A) * (B -> A)) => (fun (f : A -> B) (snd : (B -> A) * (B -> A)) => (fun (g h : B -> A) (p : c (f, (g, h)) = (idmap, idmap)) => let X0 := equiv_fun (equiv_path_prod (c (f, (g, h))) (idmap, idmap))^-1 in let p0 := X0 p in (fun (p1 : fst (c (f, (g, h))) = fst (idmap, idmap)) (q : Overture.snd (c (f, (g, h))) = Overture.snd (idmap, idmap)) => (f; ((h; ap10 p1), (g; ap10 q))) : {f0 : A -> B & {g0 : B -> A & (fun x : A => g0 (...)) == idmap} * {h0 : B -> A & (fun x : B => f0 (...)) == idmap}}) (fst p0) (Overture.snd p0)) (fst snd) (Overture.snd snd)) (fst proj1) (snd proj1)) X.1 X.2) o (fun X : {f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}} => (fun (f : A -> B) (proj2 : {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}) => (fun fst : {g : B -> A & (fun x : A => g (f x)) == idmap} => (fun (g : B -> A) (p : (fun x : A => g (f x)) == idmap) (snd : {h : B -> A & (fun x : B => f (h x)) == idmap}) => (fun (h : B -> A) (q : (fun x : B => f (h x)) == idmap) => ((f, (h, g)); path_prod (c (f, (h, g))) (idmap, idmap) (path_arrow (Overture.fst (c (f, (h, g)))) (Overture.fst (idmap, idmap)) p) (path_arrow (Overture.snd (c (f, (h, g)))) (Overture.snd (idmap, idmap)) q) : c (f, (h, g)) = (idmap, idmap))) snd.1 snd.2) fst.1 fst.2) (fst proj2) (snd proj2)) X.1 X.2) == idmap
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap

(f; ((g; ap10 (ap fst (path_prod (c (f, (h, g))) (idmap, idmap) (path_arrow (fun x : A => g (f x)) idmap p) (path_arrow (fun x : B => f (h x)) idmap q)))), (h; ap10 (ap snd (path_prod (c (f, (h, g))) (idmap, idmap) (path_arrow (fun x : A => g (f x)) idmap p) (path_arrow (fun x : B => f (h x)) idmap q)))))) = (f; ((g; p), (h; q)))
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap

ap10 (ap fst (path_prod (c (f, (h, g))) (idmap, idmap) (path_arrow (fun x : A => g (f x)) idmap p) (path_arrow (fun x : B => f (h x)) idmap q))) = p
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap
ap10 (ap snd (path_prod (c (f, (h, g))) (idmap, idmap) (path_arrow (fun x : A => g (f x)) idmap p) (path_arrow (fun x : B => f (h x)) idmap q))) = q
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap

ap10 (path_arrow (fun x : A => g (f x)) idmap p) = p
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap
ap10 (ap snd (path_prod (c (f, (h, g))) (idmap, idmap) (path_arrow (fun x : A => g (f x)) idmap p) (path_arrow (fun x : B => f (h x)) idmap q))) = q
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap

ap10 (path_arrow (fun x : A => g (f x)) idmap p) = p
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
f: A -> B
g: B -> A
p: (fun x : A => g (f x)) == idmap
h: B -> A
q: (fun x : B => f (h x)) == idmap
ap10 (path_arrow (fun x : B => f (h x)) idmap q) = q
all:apply path_forall; intros x; rewrite ap10_path_arrow; reflexivity.
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type

(fun X : {f : A -> B & {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}} => (fun (f : A -> B) (proj2 : {g : B -> A & (fun x : A => g (f x)) == idmap} * {h : B -> A & (fun x : B => f (h x)) == idmap}) => (fun fst : {g : B -> A & (fun x : A => g (f x)) == idmap} => (fun (g : B -> A) (p : (fun x : A => g (f x)) == idmap) (snd : {h : B -> A & (fun x : B => f (h x)) == idmap}) => (fun (h : B -> A) (q : (fun x : B => f (h x)) == idmap) => ((f, (h, g)); path_prod (c (f, (h, g))) (idmap, idmap) (path_arrow (Overture.fst (c (f, (h, g)))) (Overture.fst (idmap, idmap)) p) (path_arrow (Overture.snd (c (f, (h, g)))) (Overture.snd (idmap, idmap)) q) : c (f, (h, g)) = (idmap, idmap))) snd.1 snd.2) fst.1 fst.2) (fst proj2) (snd proj2)) X.1 X.2) o (fun X : U => (fun proj1 : (A -> B) * ((B -> A) * (B -> A)) => (fun (f : A -> B) (snd : (B -> A) * (B -> A)) => (fun (g h : B -> A) (p : c (f, (g, h)) = (idmap, idmap)) => let X0 := equiv_fun (equiv_path_prod (c (f, (g, h))) (idmap, idmap))^-1 in let p0 := X0 p in (fun (p1 : fst (c (f, (g, h))) = fst (idmap, idmap)) (q : Overture.snd (c (f, (g, h))) = Overture.snd (idmap, idmap)) => (f; ((h; ap10 p1), (g; ap10 q))) : {f0 : A -> B & {g0 : B -> A & (fun x : A => g0 (...)) == idmap} * {h0 : B -> A & (fun x : B => f0 (...)) == idmap}}) (fst p0) (Overture.snd p0)) (fst snd) (Overture.snd snd)) (fst proj1) (snd proj1)) X.1 X.2) == idmap
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
fghp: U

(((let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).1, ((snd (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).1, (fst (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).1)); path_prod (c ((let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).1, ((snd (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).1, (fst (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).1))) (idmap, idmap) (path_arrow (fst (c ((let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).1, ((snd (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).1, (fst (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).1)))) (fst (idmap, idmap)) (fst (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).2) (path_arrow (snd (c ((let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).1, ((snd (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).1, (fst (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).1)))) (snd (idmap, idmap)) (snd (let X := equiv_fun (equiv_path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap))^-1 in let p := X fghp.2 in (fst fghp.1; ((snd (snd fghp.1); ap10 (fst p)), (fst (snd fghp.1); ap10 (snd p))))).2).2)) = fghp
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
fghp: U

((fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1))); path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap) (path_arrow (fun x : A => snd (snd fghp.1) (fst fghp.1 x)) idmap (ap10 (ap fst fghp.2))) (path_arrow (fun x : B => fst fghp.1 (fst (snd fghp.1) x)) idmap (ap10 (ap snd fghp.2)))) = fghp
O: ReflectiveSubuniverse
H: Funext
A, B: Type
H0: In O A
H1: In O B
c:= fun X : (A -> B) * ((B -> A) * (B -> A)) => (fun x : A => snd (snd X) (fst X x), fun x : B => fst X (fst (snd X) x)): (A -> B) * ((B -> A) * (B -> A)) -> (A -> A) * (B -> B)
U:= hfiber c (idmap, idmap): Type
fghp: U

path_prod (c (fst fghp.1, (fst (snd fghp.1), snd (snd fghp.1)))) (idmap, idmap) (path_arrow (fun x : A => snd (snd fghp.1) (fst fghp.1 x)) idmap (ap10 (ap fst fghp.2))) (path_arrow (fun x : B => fst fghp.1 (fst (snd fghp.1) x)) idmap (ap10 (ap snd fghp.2))) = fghp.2
refine (_ @ eta_path_prod (pr2 fghp)); apply ap011; apply eta_path_arrow. Defined. (** ** Paths *)
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S

In O (x = y)
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S

In O (x = y)
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: O (x = y)

x = y
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: x = y
?Goal@{u:=to O (x = y) u} = u
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: O (x = y)

x = y
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: O (x = y)

(fun _ : O (x = y) => x) == (fun _ : O (x = y) => y)
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: O (x = y)
p: (fun _ : O (x = y) => x) == (fun _ : O (x = y) => y)
x = y
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: O (x = y)

(fun _ : O (x = y) => x) == (fun _ : O (x = y) => y)
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: O (x = y)

(fun _ : x = y => x) == (fun _ : x = y => y)
intro v; exact v.
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: O (x = y)
p: (fun _ : O (x = y) => x) == (fun _ : O (x = y) => y)

x = y
exact (p u).
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: x = y

(let p := O_indpaths (fun _ : O (x = y) => x) (fun _ : O (x = y) => y) ((idmap : (fun _ : x = y => x) == (fun _ : x = y => y)) : (fun _ : x = y => x) == (fun _ : x = y => y)) in p (to O (x = y) u)) = u
O: ReflectiveSubuniverse
S: Type
S_inO: In O S
x, y: S
u: x = y

O_indpaths (fun _ : O (x = y) => x) (fun _ : O (x = y) => y) idmap (to O (x = y) u) = u
rewrite O_indpaths_beta; reflexivity. Qed. Global Existing Instance inO_paths.
O: ReflectiveSubuniverse
A: Type
a0, a1, a2: A

O (a0 = a1) -> O (a1 = a2) -> O (a0 = a2)
O: ReflectiveSubuniverse
A: Type
a0, a1, a2: A

O (a0 = a1) -> O (a1 = a2) -> O (a0 = a2)
O: ReflectiveSubuniverse
A: Type
a0, a1, a2: A
p: O (a0 = a1)
q: O (a1 = a2)

O (a0 = a2)
O: ReflectiveSubuniverse
A: Type
a0, a1, a2: A
q: a1 = a2
p: a0 = a1

O (a0 = a2)
exact (to O _ (p @ q)). Defined. (** ** Truncations *) (** The reflector preserves hprops (and, as we have already seen, contractible types), although it doesn't generally preserve [n]-types for other [n]. *)
O: ReflectiveSubuniverse
A: Type
IsHProp0: IsHProp A

IsHProp (O A)
O: ReflectiveSubuniverse
A: Type
IsHProp0: IsHProp A

IsHProp (O A)
O: ReflectiveSubuniverse
A: Type
IsHProp0: IsHProp A

IsEquiv (fun a : O A => (a, a))
O: ReflectiveSubuniverse
A: Type
IsHProp0: IsHProp A

(fun x : O A => O_prod_cmp A A (O_functor (fun a : A => (a, a)) x)) == (fun a : O A => (a, a))
O: ReflectiveSubuniverse
A: Type
IsHProp0: IsHProp A
x: A

O_prod_cmp A A (O_functor (fun a : A => (a, a)) (to O A x)) = (to O A x, to O A x)
O: ReflectiveSubuniverse
A: Type
IsHProp0: IsHProp A
x: A

O_prod_cmp A A (to O (A * A) (x, x)) = (to O A x, to O A x)
unfold O_prod_cmp; apply O_rec_beta. Defined. (** If [A] is [In O], then so is [IsTrunc n A]. *)
O: ReflectiveSubuniverse
H: Funext
n: trunc_index
A: Type
H0: In O A

In O (IsTrunc n A)
O: ReflectiveSubuniverse
H: Funext
n: trunc_index
A: Type
H0: In O A

In O (IsTrunc n A)
O: ReflectiveSubuniverse
H: Funext
A: Type
H0: In O A

In O (Contr A)
O: ReflectiveSubuniverse
H: Funext
n: trunc_index
IH: forall A : Type, In O A -> In O (IsTrunc n A)
A: Type
H0: In O A
In O (IsTrunc n.+1 A)
O: ReflectiveSubuniverse
H: Funext
A: Type
H0: In O A

In O (Contr A)
refine (inO_equiv_inO _ equiv_contr_inhabited_allpath^-1).
O: ReflectiveSubuniverse
H: Funext
n: trunc_index
IH: forall A : Type, In O A -> In O (IsTrunc n A)
A: Type
H0: In O A

In O (IsTrunc n.+1 A)
refine (inO_equiv_inO _ (equiv_istrunc_unfold n.+1 A)^-1). Defined. (** ** Coproducts *)
O: ReflectiveSubuniverse
A, B, A', B': Type
f: A -> A'
g: B -> B'
O_inverts0: O_inverts f
O_inverts1: O_inverts g

O_inverts (functor_sum f g)
O: ReflectiveSubuniverse
A, B, A', B': Type
f: A -> A'
g: B -> B'
O_inverts0: O_inverts f
O_inverts1: O_inverts g

O_inverts (functor_sum f g)
O: ReflectiveSubuniverse
A, B, A', B': Type
f: A -> A'
g: B -> B'
O_inverts0: O_inverts f
O_inverts1: O_inverts g
Z: Type
X: In O Z

ExtendableAlong 2 (functor_sum f g) (fun _ : A' + B' => Z)
apply extendable_functor_sum; apply ooextendable_O_inverts; assumption. Defined. Definition equiv_O_functor_sum {A B A' B'} (f : A -> A') (g : B -> B') `{O_inverts f} `{O_inverts g} : O (A + B) <~> O (A' + B') := Build_Equiv _ _ _ (O_inverts_functor_sum f g). Definition equiv_O_sum {A B} : O (A + B) <~> O (O A + O B) := equiv_O_functor_sum (to O A) (to O B). (** ** Coequalizers *) Section OCoeq. Context {B A : Type} (f g : B -> A).
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts k
O_inverts1: O_inverts h

O_inverts (functor_coeq h k p q)
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts k
O_inverts1: O_inverts h

O_inverts (functor_coeq h k p q)
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts k
O_inverts1: O_inverts h

forall Z : Type, In O Z -> ExtendableAlong 2 (functor_coeq h k p q) (fun _ : Coeq f' g' => Z)
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts k
O_inverts1: O_inverts h
Z: Type
Z_inO: In O Z

ExtendableAlong 2 (functor_coeq h k p q) (fun _ : Coeq f' g' => Z)
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts k
O_inverts1: O_inverts h
Z: Type
Z_inO: In O Z

ExtendableAlong 2 k (fun _ : A' => Z)
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts k
O_inverts1: O_inverts h
Z: Type
Z_inO: In O Z
ExtendableAlong 3 h (fun _ : B' => Z)
all:nrapply ooextendable_O_inverts; assumption. Defined. Definition equiv_O_functor_coeq {B' A' : Type} (f' g' : B' -> A') (h : B -> B') (k : A -> A') (p : k o f == f' o h) (q : k o g == g' o h) `{O_inverts k} `{O_inverts h} : O (Coeq f g) <~> O (Coeq f' g') := Build_Equiv _ _ _ (O_inverts_functor_coeq h k p q). Definition coeq_cmp : Coeq f g -> Coeq (O_functor f) (O_functor g) := functor_coeq (to O B) (to O A) (fun y => (to_O_natural f y)^) (fun y => (to_O_natural g y)^).
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A

O_inverts coeq_cmp
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A

O_inverts coeq_cmp
rapply O_inverts_functor_coeq. Defined. Definition equiv_O_coeq : O (Coeq f g) <~> O (Coeq (O_functor f) (O_functor g)) := Build_Equiv _ _ (O_functor coeq_cmp) _.
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
a: A

equiv_O_coeq (to O (Coeq f g) (coeq a)) = to O (Coeq (O_functor f) (O_functor g)) (coeq (to O A a))
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
a: A

equiv_O_coeq (to O (Coeq f g) (coeq a)) = to O (Coeq (O_functor f) (O_functor g)) (coeq (to O A a))
refine (to_O_natural _ _). Defined.
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
a: A

(O_functor coeq_cmp)^-1 (to O (Coeq (O_functor f) (O_functor g)) (coeq (to O A a))) = to O (Coeq f g) (coeq a)
O: ReflectiveSubuniverse
B, A: Type
f, g: B -> A
a: A

(O_functor coeq_cmp)^-1 (to O (Coeq (O_functor f) (O_functor g)) (coeq (to O A a))) = to O (Coeq f g) (coeq a)
apply moveR_equiv_V; symmetry; apply equiv_O_coeq_to_O. Defined. End OCoeq. (** ** Pushouts *) Section OPushout. Context {A B C : Type} (f : A -> B) (g : A -> C).
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
A', B', C': Type
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
O_inverts0: O_inverts h
O_inverts1: O_inverts k
O_inverts2: O_inverts l

O_inverts (functor_pushout h k l p q)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
A', B', C': Type
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
O_inverts0: O_inverts h
O_inverts1: O_inverts k
O_inverts2: O_inverts l

O_inverts (functor_pushout h k l p q)
rapply O_inverts_functor_coeq; rapply O_inverts_functor_sum. Defined. Definition equiv_O_pushout : O (Pushout f g) <~> O (Pushout (O_functor f) (O_functor g)) := Build_Equiv _ _ _ (O_inverts_functor_pushout (to O A) (to O B) (to O C) (fun x => (to_O_natural f x)^) (fun x => (to_O_natural g x)^)).
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
b: B

equiv_O_pushout (to O (Pushout f g) (pushl b)) = to O (Pushout (O_functor f) (O_functor g)) (pushl (to O B b))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
b: B

equiv_O_pushout (to O (Pushout f g) (pushl b)) = to O (Pushout (O_functor f) (O_functor g)) (pushl (to O B b))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
b: B

O_functor (functor_pushout (to O A) (to O B) (to O C) (fun x : A => (to_O_natural f x)^) (fun x : A => (to_O_natural g x)^)) (to O (Pushout f g) (pushl b)) = to O (Pushout (O_functor f) (O_functor g)) (pushl (to O B b))
rapply to_O_natural. Defined.
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
c: C

equiv_O_pushout (to O (Pushout f g) (pushr c)) = to O (Pushout (O_functor f) (O_functor g)) (pushr (to O C c))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
c: C

equiv_O_pushout (to O (Pushout f g) (pushr c)) = to O (Pushout (O_functor f) (O_functor g)) (pushr (to O C c))
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
c: C

O_functor (functor_pushout (to O A) (to O B) (to O C) (fun x : A => (to_O_natural f x)^) (fun x : A => (to_O_natural g x)^)) (to O (Pushout f g) (pushr c)) = to O (Pushout (O_functor f) (O_functor g)) (pushr (to O C c))
rapply to_O_natural. Defined.
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
b: B

equiv_O_pushout^-1 (to O (Pushout (O_functor f) (O_functor g)) (pushl (to O B b))) = to O (Pushout f g) (pushl b)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
b: B

equiv_O_pushout^-1 (to O (Pushout (O_functor f) (O_functor g)) (pushl (to O B b))) = to O (Pushout f g) (pushl b)
apply moveR_equiv_V; symmetry; apply equiv_O_pushout_to_O_pushl. Qed.
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
c: C

equiv_O_pushout^-1 (to O (Pushout (O_functor f) (O_functor g)) (pushr (to O C c))) = to O (Pushout f g) (pushr c)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: A -> C
c: C

equiv_O_pushout^-1 (to O (Pushout (O_functor f) (O_functor g)) (pushr (to O C c))) = to O (Pushout f g) (pushr c)
apply moveR_equiv_V; symmetry; apply equiv_O_pushout_to_O_pushr. Qed. End OPushout. End Types. Section Decidable. (** If [Empty] belongs to [O], then [O] preserves decidability. *)
O: ReflectiveSubuniverse
H: In O Empty
A: Type
H0: Decidable A

Decidable (O A)
O: ReflectiveSubuniverse
H: In O Empty
A: Type
H0: Decidable A

Decidable (O A)
O: ReflectiveSubuniverse
H: In O Empty
A: Type
H0: Decidable A
y: A

Decidable (O A)
O: ReflectiveSubuniverse
H: In O Empty
A: Type
H0: Decidable A
n: ~ A
Decidable (O A)
O: ReflectiveSubuniverse
H: In O Empty
A: Type
H0: Decidable A
y: A

Decidable (O A)
exact (inl (to O A y)).
O: ReflectiveSubuniverse
H: In O Empty
A: Type
H0: Decidable A
n: ~ A

Decidable (O A)
exact (inr (O_rec n)). Defined. (** Dually, if [O A] is decidable, then [O (Decidable A)]. *)
O: ReflectiveSubuniverse
A: Type
H: Decidable (O A)

O (Decidable A)
O: ReflectiveSubuniverse
A: Type
H: Decidable (O A)

O (Decidable A)
O: ReflectiveSubuniverse
A: Type
H: Decidable (O A)
y: O A

O (Decidable A)
O: ReflectiveSubuniverse
A: Type
H: Decidable (O A)
n: ~ O A
O (Decidable A)
O: ReflectiveSubuniverse
A: Type
H: Decidable (O A)
y: O A

O (Decidable A)
exact (O_functor inl y).
O: ReflectiveSubuniverse
A: Type
H: Decidable (O A)
n: ~ O A

O (Decidable A)
O: ReflectiveSubuniverse
A: Type
H: Decidable (O A)
n: ~ O A

O (~ A)
O: ReflectiveSubuniverse
A: Type
H: Decidable (O A)
n: ~ O A
a: A

Empty
exact (n (to O A a)). Defined. End Decidable. Section Monad. Definition O_monad_mult (A : Type) : O (O A) -> O A := O_rec idmap.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B

O_functor f o O_monad_mult A == O_monad_mult B o O_functor (O_functor f)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B

O_functor f o O_monad_mult A == O_monad_mult B o O_functor (O_functor f)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
x: O A

O_functor f (O_rec idmap (to O (O A) x)) = O_rec idmap (O_functor (O_functor f) (to O (O A) x))
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
x: O A

O_functor f (O_rec idmap (to O (O A) x)) = O_rec idmap (to O (O B) (O_functor f x))
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
x: O A

O_functor f x = O_rec idmap (to O (O B) (O_functor f x))
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
x: O A

O_functor f x = O_functor f x
reflexivity. Qed.
O: ReflectiveSubuniverse
A: Type

O_monad_mult A o to O (O A) == idmap
O: ReflectiveSubuniverse
A: Type

O_monad_mult A o to O (O A) == idmap
O: ReflectiveSubuniverse
A: Type
x: A

O_rec idmap (to O (O A) (to O A x)) = to O A x
exact (O_rec_beta idmap (to O A x)). Defined.
O: ReflectiveSubuniverse
A: Type

O_monad_mult A o O_functor (to O A) == idmap
O: ReflectiveSubuniverse
A: Type

O_monad_mult A o O_functor (to O A) == idmap
O: ReflectiveSubuniverse
A: Type
x: A

O_rec idmap (O_rec (fun x : A => to O (O_reflector O A) (to O A x)) (to O A x)) = to O A x
O: ReflectiveSubuniverse
A: Type
x: A

to O A x = to O A x
reflexivity. Qed.
O: ReflectiveSubuniverse
A: Type

O_monad_mult A o O_monad_mult (O A) == O_monad_mult A o O_functor (O_monad_mult A)
O: ReflectiveSubuniverse
A: Type

O_monad_mult A o O_monad_mult (O A) == O_monad_mult A o O_functor (O_monad_mult A)
O: ReflectiveSubuniverse
A: Type
x: O (O A)

O_rec idmap (O_rec idmap (to O (O (O A)) x)) = O_rec idmap (O_rec (fun x : O (O A) => to O (O A) (O_rec idmap x)) (to O (O (O A)) x))
O: ReflectiveSubuniverse
A: Type
x: O (O A)

O_rec idmap x = O_rec idmap x
reflexivity. Qed. End Monad. Section StrongMonad. Context {fs : Funext}. Definition O_monad_strength (A B : Type) : A * O B -> O (A * B) := fun aob => O_rec (fun b a => to O (A*B) (a,b)) (snd aob) (fst aob).
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'

O_functor (functor_prod f g) o O_monad_strength A B == O_monad_strength A' B' o functor_prod f (O_functor g)
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'

O_functor (functor_prod f g) o O_monad_strength A B == O_monad_strength A' B' o functor_prod f (O_functor g)
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'
a: A
b: O B

O_functor (functor_prod f g) (O_monad_strength A B (a, b)) = O_monad_strength A' B' (functor_prod f (O_functor g) (a, b))
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'
b: O B

forall a : A, O_functor (functor_prod f g) (O_monad_strength A B (a, b)) = O_monad_strength A' B' (functor_prod f (O_functor g) (a, b))
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'
b: O B

(fun x : A => O_functor (functor_prod f g) (O_monad_strength A B (x, b))) = (fun x : A => O_monad_strength A' B' (functor_prod f (O_functor g) (x, b)))
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'
b: B

(fun x : A => O_functor (functor_prod f g) (O_monad_strength A B (x, to O B b))) = (fun x : A => O_monad_strength A' B' (functor_prod f (O_functor g) (x, to O B b)))
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'
b: B
a: A

O_functor (functor_prod f g) (O_monad_strength A B (a, to O B b)) = O_monad_strength A' B' (functor_prod f (O_functor g) (a, to O B b))
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'
b: B
a: A

O_rec (fun x : A * B => to O (A' * B') (functor_prod f g x)) (O_rec (fun (b : B) (a : A) => to O (A * B) (a, b)) (to O B b) a) = O_rec (fun (b : B') (a : A') => to O (A' * B') (a, b)) (O_rec (fun x : B => to O B' (g x)) (to O B b)) (f a)
O: ReflectiveSubuniverse
fs: Funext
A, A', B, B': Type
f: A -> A'
g: B -> B'
b: B
a: A

to O (A' * B') (functor_prod f g (a, b)) = to O (A' * B') (f a, g b)
reflexivity. Qed. (** The diagrams for strength, see http://en.wikipedia.org/wiki/Strong_monad *)
O: ReflectiveSubuniverse
fs: Funext
A: Type

O_functor snd o O_monad_strength Unit A == snd
O: ReflectiveSubuniverse
fs: Funext
A: Type

O_functor snd o O_monad_strength Unit A == snd
O: ReflectiveSubuniverse
fs: Funext
A: Type
a: O A

O_functor snd (O_monad_strength Unit A (tt, a)) = snd (tt, a)
O: ReflectiveSubuniverse
fs: Funext
A: Type
a: A

O_functor snd (O_monad_strength Unit A (tt, to O A a)) = snd (tt, to O A a)
O: ReflectiveSubuniverse
fs: Funext
A: Type
a: A

O_rec (fun x : Unit * A => to O A (snd x)) (O_rec (fun (b : A) (a : Unit) => to O (Unit * A) (a, b)) (snd (tt, to O A a)) (fst (tt, to O A a))) = snd (tt, to O A a)
O: ReflectiveSubuniverse
fs: Funext
A: Type
a: A

O_rec (fun x : Unit * A => to O A (snd x)) (O_rec (fun (b : A) (a : Unit) => to O (Unit * A) (a, b)) (to O A a) tt) = to O A a
O: ReflectiveSubuniverse
fs: Funext
A: Type
a: A

O_rec (fun x : Unit * A => to O A (snd x)) (to O (Unit * A) (tt, a)) = to O A a
nrapply O_rec_beta. Qed.
O: ReflectiveSubuniverse
fs: Funext
A, B: Type

O_monad_strength A B o functor_prod idmap (to O B) == to O (A * B)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type

O_monad_strength A B o functor_prod idmap (to O B) == to O (A * B)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
a: A
b: B

O_monad_strength A B (functor_prod idmap (to O B) (a, b)) = to O (A * B) (a, b)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
a: A
b: B

O_rec (fun (b : B) (a : A) => to O (A * B) (a, b)) (snd (fst (a, b), to O B (snd (a, b)))) (fst (fst (a, b), to O B (snd (a, b)))) = to O (A * B) (a, b)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
a: A
b: B

O_rec (fun (b : B) (a : A) => to O (A * B) (a, b)) (to O B b) a = to O (A * B) (a, b)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
b: B

O_rec (fun (b : B) (a : A) => to O (A * B) (a, b)) (to O B b) = (fun x : A => to O (A * B) (x, b))
nrapply O_rec_beta. Qed.
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type

O_functor (equiv_prod_assoc A B C)^-1 o O_monad_strength (A * B) C == O_monad_strength A (B * C) o functor_prod idmap (O_monad_strength B C) o (equiv_prod_assoc A B (O C))^-1
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type

O_functor (equiv_prod_assoc A B C)^-1 o O_monad_strength (A * B) C == O_monad_strength A (B * C) o functor_prod idmap (O_monad_strength B C) o (equiv_prod_assoc A B (O C))^-1
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
a: A
b: B
c: O C

O_functor (equiv_prod_assoc A B C)^-1 (O_monad_strength (A * B) C (a, b, c)) = O_monad_strength A (B * C) (functor_prod idmap (O_monad_strength B C) ((equiv_prod_assoc A B (O C))^-1 (a, b, c)))
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
b: B
c: O C

(fun x : A => O_functor (equiv_prod_assoc A B C)^-1 (O_monad_strength (A * B) C (x, b, c))) = (fun x : A => O_monad_strength A (B * C) (functor_prod idmap (O_monad_strength B C) ((equiv_prod_assoc A B (O C))^-1 (x, b, c))))
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
c: O C

(fun (x : B) (x0 : A) => O_functor (equiv_prod_assoc A B C)^-1 (O_monad_strength (A * B) C (x0, x, c))) = (fun (x : B) (x0 : A) => O_monad_strength A (B * C) (functor_prod idmap (O_monad_strength B C) ((equiv_prod_assoc A B (O C))^-1 (x0, x, c))))
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
c: C

(fun (x : B) (x0 : A) => O_functor (equiv_prod_assoc A B C)^-1 (O_monad_strength (A * B) C (x0, x, to O C c))) = (fun (x : B) (x0 : A) => O_monad_strength A (B * C) (functor_prod idmap (O_monad_strength B C) ((equiv_prod_assoc A B (O C))^-1 (x0, x, to O C c))))
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
c: C
b: B

(fun x : A => O_functor (equiv_prod_assoc A B C)^-1 (O_monad_strength (A * B) C (x, b, to O C c))) = (fun x : A => O_monad_strength A (B * C) (functor_prod idmap (O_monad_strength B C) ((equiv_prod_assoc A B (O C))^-1 (x, b, to O C c))))
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
c: C
b: B
a: A

O_functor (equiv_prod_assoc A B C)^-1 (O_monad_strength (A * B) C (a, b, to O C c)) = O_monad_strength A (B * C) (functor_prod idmap (O_monad_strength B C) ((equiv_prod_assoc A B (O C))^-1 (a, b, to O C c)))
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
c: C
b: B
a: A

O_rec (fun x : A * B * C => to O (A * (B * C)) ((equiv_prod_assoc A B C)^-1 x)) (O_rec (fun (b : C) (a : A * B) => to O (A * B * C) (a, b)) (snd (a, b, to O C c)) (fst (a, b, to O C c))) = O_rec (fun (b : B * C) (a : A) => to O (A * (B * C)) (a, b)) (snd (fst ((equiv_prod_assoc A B (O C))^-1 (a, b, to O C c)), O_rec (fun (b : C) (a : B) => to O (B * C) (a, b)) (snd (snd ((equiv_prod_assoc A B (O C))^-1 (a, b, to O C c)))) (fst (snd ((equiv_prod_assoc A B (O C))^-1 (a, b, to O C c)))))) (fst (fst ((equiv_prod_assoc A B (O C))^-1 (a, b, to O C c)), O_rec (fun (b : C) (a : B) => to O (B * C) (a, b)) (snd (snd ((equiv_prod_assoc A B (O C))^-1 (a, b, to O C c)))) (fst (snd ((equiv_prod_assoc A B (O C))^-1 (a, b, to O C c))))))
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
c: C
b: B
a: A

O_rec (fun x : A * B * C => to O (A * (B * C)) (fst (fst x), (snd (fst x), snd x))) (O_rec (fun (b : C) (a : A * B) => to O (A * B * C) (a, b)) (to O C c) (a, b)) = O_rec (fun (b : B * C) (a : A) => to O (A * (B * C)) (a, b)) (O_rec (fun (b : C) (a : B) => to O (B * C) (a, b)) (to O C c) b) a
O: ReflectiveSubuniverse
fs: Funext
A, B, C: Type
c: C
b: B
a: A

to O (A * (B * C)) (a, (b, c)) = to O (A * (B * C)) (a, (b, c))
reflexivity. Qed.
O: ReflectiveSubuniverse
fs: Funext
A, B: Type

O_monad_mult (A * B) o O_functor (O_monad_strength A B) o O_monad_strength A (O B) == O_monad_strength A B o functor_prod idmap (O_monad_mult B)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type

O_monad_mult (A * B) o O_functor (O_monad_strength A B) o O_monad_strength A (O B) == O_monad_strength A B o functor_prod idmap (O_monad_mult B)
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
a: A
b: O (O B)

O_monad_mult (A * B) (O_functor (O_monad_strength A B) (O_monad_strength A (O B) (a, b))) = O_monad_strength A B (functor_prod idmap (O_monad_mult B) (a, b))
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
b: O (O B)

(fun x : A => O_monad_mult (A * B) (O_functor (O_monad_strength A B) (O_monad_strength A (O B) (x, b)))) = (fun x : A => O_monad_strength A B (functor_prod idmap (O_monad_mult B) (x, b)))
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
b: B

(fun x : A => O_monad_mult (A * B) (O_functor (O_monad_strength A B) (O_monad_strength A (O B) (x, to O (O B) (to O B b))))) = (fun x : A => O_monad_strength A B (functor_prod idmap (O_monad_mult B) (x, to O (O B) (to O B b))))
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
b: B
a: A

O_monad_mult (A * B) (O_functor (O_monad_strength A B) (O_monad_strength A (O B) (a, to O (O B) (to O B b)))) = O_monad_strength A B (functor_prod idmap (O_monad_mult B) (a, to O (O B) (to O B b)))
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
b: B
a: A

O_rec idmap (O_rec (fun x : A * O B => to O (O (A * B)) (O_rec (fun (b : B) (a : A) => to O (A * B) (a, b)) (snd x) (fst x))) (O_rec (fun (b : O B) (a : A) => to O (A * O B) (a, b)) (snd (a, to O (O B) (to O B b))) (fst (a, to O (O B) (to O B b))))) = O_rec (fun (b : B) (a : A) => to O (A * B) (a, b)) (snd (fst (a, to O (O B) (to O B b)), O_rec idmap (snd (a, to O (O B) (to O B b))))) (fst (fst (a, to O (O B) (to O B b)), O_rec idmap (snd (a, to O (O B) (to O B b)))))
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
b: B
a: A

O_rec idmap (O_rec (fun x : A * O B => to O (O (A * B)) (O_rec (fun (b : B) (a : A) => to O (A * B) (a, b)) (snd x) (fst x))) (O_rec (fun (b : O B) (a : A) => to O (A * O B) (a, b)) (to O (O B) (to O B b)) a)) = O_rec (fun (b : B) (a : A) => to O (A * B) (a, b)) (O_rec idmap (to O (O B) (to O B b))) a
O: ReflectiveSubuniverse
fs: Funext
A, B: Type
b: B
a: A

to O (A * B) (a, b) = to O (A * B) (a, b)
reflexivity. Qed. End StrongMonad. End Reflective_Subuniverse. (** Now we make the [O_inverts] notation global. *) Notation O_inverts O f := (IsEquiv (O_functor O f)). (** ** Modally connected types *) (** Connectedness of a type, relative to a modality or reflective subuniverse, can be defined in two equivalent ways: quantifying over all maps into modal types, or by considering just the universal case, the modal reflection of the type itself. The former requires only core Coq, but blows up the size (universe level) of [IsConnected], since it quantifies over types; moreover, it is not even quite correct since (at least with a polymorphic modality) it should really be quantified over all universes. Thus, we use the latter, although in most examples it requires HITs to define the modal reflection. Question: is there a definition of connectedness (say, for n-types) that neither blows up the universe level, nor requires HIT's? *) (** We give annotations to reduce the number of universe parameters. *) Class IsConnected (O : ReflectiveSubuniverse@{i}) (A : Type@{i}) := isconnected_contr_O : Contr@{i} (O A). Global Existing Instance isconnected_contr_O. Section ConnectedTypes. Context (O : ReflectiveSubuniverse). (** Being connected is an hprop *)
O: ReflectiveSubuniverse
H: Funext
A: Type

IsHProp (IsConnected O A)
O: ReflectiveSubuniverse
H: Funext
A: Type

IsHProp (IsConnected O A)
unfold IsConnected; exact _. Defined. (** Anything equivalent to a connected type is connected. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

IsConnected O A -> IsConnected O B
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

IsConnected O A -> IsConnected O B
intros ?; refine (contr_equiv (O A) (O_functor O f)). Defined. Definition isconnected_equiv' (A : Type) {B : Type} (f : A <~> B) : IsConnected O A -> IsConnected O B := isconnected_equiv A f. (** The O-connected types form a subuniverse. *)
O: ReflectiveSubuniverse

Subuniverse
O: ReflectiveSubuniverse

Subuniverse
O: ReflectiveSubuniverse

forall T U : Type, IsConnected O T -> forall f : T -> U, IsEquiv f -> IsConnected O U
O: ReflectiveSubuniverse
T, U: Type
isconnT: IsConnected O T
f: T -> U
isequivf: IsEquiv f

IsConnected O U
exact (isconnected_equiv T f isconnT). Defined. (** Connectedness of a type [A] can equivalently be characterized by the fact that any map to an [O]-type [C] is nullhomotopic. Here is one direction of that equivalence. *)
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
C: Type
H0: In O C
f: A -> C

NullHomotopy f
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
C: Type
H0: In O C
f: A -> C

NullHomotopy f
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
C: Type
H0: In O C
f: A -> C
ff:= O_rec f: O_reflector O A -> C

NullHomotopy f
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
C: Type
H0: In O C
f: A -> C
ff:= O_rec f: O_reflector O A -> C

forall x : A, f x = ff (center (O_reflector O A))
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
C: Type
H0: In O C
f: A -> C
ff:= O_rec f: O_reflector O A -> C
a: A

f a = ff (center (O_reflector O A))
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
C: Type
H0: In O C
f: A -> C
ff:= O_rec f: O_reflector O A -> C
a: A

ff (center (O_reflector O A)) = f a
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
C: Type
H0: In O C
f: A -> C
ff:= O_rec f: O_reflector O A -> C
a: A

ff (to O A a) = f a
apply O_rec_beta. Defined. (** For the other direction of the equivalence, it's sufficient to consider the case when [C] is [O A]. *)
O: ReflectiveSubuniverse
A: Type

NullHomotopy (to O A) -> IsConnected O A
O: ReflectiveSubuniverse
A: Type

NullHomotopy (to O A) -> IsConnected O A
O: ReflectiveSubuniverse
A: Type
nh: NullHomotopy (to O A)

IsConnected O A
O: ReflectiveSubuniverse
A: Type
nh: NullHomotopy (to O A)

forall y : O_reflector O A, nh.1 = y
O: ReflectiveSubuniverse
A: Type
nh: NullHomotopy (to O A)

(fun _ : A => nh.1) == (fun x : A => to O A x)
intros x; symmetry; apply (nh .2). Defined. (** Now the general case follows. *)
O: ReflectiveSubuniverse
A: Type

(forall C : Type, In O C -> forall f : A -> C, NullHomotopy f) -> IsConnected O A
O: ReflectiveSubuniverse
A: Type

(forall C : Type, In O C -> forall f : A -> C, NullHomotopy f) -> IsConnected O A
O: ReflectiveSubuniverse
A: Type
H: forall C : Type, In O C -> forall f : A -> C, NullHomotopy f

IsConnected O A
exact (isconnected_from_elim_to_O (H (O A) (O_inO A) (to O A))). Defined. (** Connected types are closed under sigmas. *)
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: IsConnected O A
H0: forall a : A, IsConnected O (B a)

IsConnected O {a : A & B a}
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: IsConnected O A
H0: forall a : A, IsConnected O (B a)

IsConnected O {a : A & B a}
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: IsConnected O A
H0: forall a : A, IsConnected O (B a)
C: Type
H1: In O C
f: {a : A & B a} -> C

NullHomotopy f
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: IsConnected O A
H0: forall a : A, IsConnected O (B a)
C: Type
H1: In O C
f: {a : A & B a} -> C
nB:= fun a : A => isconnected_elim C (fun b : B a => f (a; b)): forall a : A, NullHomotopy (fun b : B a => f (a; b))

NullHomotopy f
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: IsConnected O A
H0: forall a : A, IsConnected O (B a)
C: Type
H1: In O C
f: {a : A & B a} -> C
nB:= fun a : A => isconnected_elim C (fun b : B a => f (a; b)): forall a : A, NullHomotopy (fun b : B a => f (a; b))
nA:= isconnected_elim C (fun a : A => (nB a).1): NullHomotopy (fun a : A => (nB a).1)

NullHomotopy f
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: IsConnected O A
H0: forall a : A, IsConnected O (B a)
C: Type
H1: In O C
f: {a : A & B a} -> C
nB:= fun a : A => isconnected_elim C (fun b : B a => f (a; b)): forall a : A, NullHomotopy (fun b : B a => f (a; b))
nA:= isconnected_elim C (fun a : A => (nB a).1): NullHomotopy (fun a : A => (nB a).1)
a: A
b: B a

f (a; b) = nA.1
exact ((nB a).2 b @ nA.2 a). Defined. (** Contractible types are connected. *)
O: ReflectiveSubuniverse
A: Type
Contr0: Contr A

IsConnected O A
O: ReflectiveSubuniverse
A: Type
Contr0: Contr A

IsConnected O A
rapply contr_O_contr. Defined. (** A type which is both connected and modal is contractible. *)
O: ReflectiveSubuniverse
A: Type
H: In O A
H0: IsConnected O A

Contr A
O: ReflectiveSubuniverse
A: Type
H: In O A
H0: IsConnected O A

Contr A
apply (contr_equiv _ (to O A)^-1). Defined. (** Any map between connected types is inverted by O. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnected O A
H0: IsConnected O B

O_inverts O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnected O A
H0: IsConnected O B

O_inverts O f
exact _. Defined. (** Here's another way of stating the universal property for mapping out of connected types into modal ones. *)
O: ReflectiveSubuniverse
n: nat
A: Type
H: IsConnected O A
C: Type
H0: In O C

ExtendableAlong n (const_tt A) (unit_name C)
O: ReflectiveSubuniverse
n: nat
A: Type
H: IsConnected O A
C: Type
H0: In O C

ExtendableAlong n (const_tt A) (unit_name C)
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
n: nat
IHn: forall C : Type, In O C -> ExtendableAlong n (const_tt A) (unit_name C)
C: Type
H0: In O C

forall g : A -> C, ExtensionAlong (const_tt A) (unit_name C) g
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
n: nat
IHn: forall C : Type, In O C -> ExtendableAlong n (const_tt A) (unit_name C)
C: Type
H0: In O C
forall h k : Unit -> C, ExtendableAlong n (const_tt A) (fun b : Unit => h b = k b)
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
n: nat
IHn: forall C : Type, In O C -> ExtendableAlong n (const_tt A) (unit_name C)
C: Type
H0: In O C

forall g : A -> C, ExtensionAlong (const_tt A) (unit_name C) g
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
n: nat
IHn: forall C : Type, In O C -> ExtendableAlong n (const_tt A) (unit_name C)
C: Type
H0: In O C
f: A -> C

ExtensionAlong (const_tt A) (unit_name C) f
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
n: nat
IHn: forall C : Type, In O C -> ExtendableAlong n (const_tt A) (unit_name C)
C: Type
H0: In O C
f: A -> C
a: A

(isconnected_elim C f).1 = f a
symmetry; apply ((isconnected_elim C f).2).
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
n: nat
IHn: forall C : Type, In O C -> ExtendableAlong n (const_tt A) (unit_name C)
C: Type
H0: In O C

forall h k : Unit -> C, ExtendableAlong n (const_tt A) (fun b : Unit => h b = k b)
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
n: nat
IHn: forall C : Type, In O C -> ExtendableAlong n (const_tt A) (unit_name C)
C: Type
H0: In O C
h, k: Unit -> C

ExtendableAlong n (const_tt A) (fun b : Unit => h b = k b)
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
n: nat
IHn: forall C : Type, In O C -> ExtendableAlong n (const_tt A) (unit_name C)
C: Type
H0: In O C
h, k: Unit -> C

forall b : Unit, h tt = k tt <~> h b = k b
intros []; apply equiv_idmap. Defined. Definition ooextendable_const_isconnected_inO (A : Type@{i}) `{IsConnected@{i} O A} (C : Type@{j}) `{In O C} : ooExtendableAlong (const_tt A) (fun _ => C) := fun n => extendable_const_isconnected_inO n A C.
O: ReflectiveSubuniverse
H: Funext
A: Type
H0: IsConnected O A
C: Type
H1: In O C

IsEquiv const
O: ReflectiveSubuniverse
H: Funext
A: Type
H0: IsConnected O A
C: Type
H1: In O C

IsEquiv const
refine (@isequiv_compose _ _ (fun c u => c) _ _ _ (isequiv_ooextendable (fun _ => C) (const_tt A) (ooextendable_const_isconnected_inO A C))). Defined. Definition equiv_const_isconnected_inO `{Funext} {A : Type} `{IsConnected O A} (C : Type) `{In O C} : C <~> (A -> C) := Build_Equiv _ _ const (isequiv_const_isconnected_inO C). End ConnectedTypes. (** ** Modally truncated maps *) Section ModalMaps. Context (O : ReflectiveSubuniverse). (** Any equivalence is modal *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

MapIn O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

MapIn O f
intros b; exact _. Defined. (** A slightly specialized result: if [Empty] is modal, then a map with decidable hprop fibers (such as [inl] or [inr]) is modal. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O Empty
H0: forall b : B, IsHProp (hfiber f b)
H1: forall b : B, Decidable (hfiber f b)

MapIn O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O Empty
H0: forall b : B, IsHProp (hfiber f b)
H1: forall b : B, Decidable (hfiber f b)

MapIn O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O Empty
H0: forall b : B, IsHProp (hfiber f b)
H1: forall b : B, Decidable (hfiber f b)
b: B

In O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O Empty
H0: forall b : B, IsHProp (hfiber f b)
H1: forall b : B, Decidable (hfiber f b)
b: B
e: hfiber f b <~> Unit

In O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O Empty
H0: forall b : B, IsHProp (hfiber f b)
H1: forall b : B, Decidable (hfiber f b)
b: B
e: hfiber f b <~> Empty
In O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O Empty
H0: forall b : B, IsHProp (hfiber f b)
H1: forall b : B, Decidable (hfiber f b)
b: B
e: hfiber f b <~> Unit

In O (hfiber f b)
exact (inO_equiv_inO Unit e^-1).
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O Empty
H0: forall b : B, IsHProp (hfiber f b)
H1: forall b : B, Decidable (hfiber f b)
b: B
e: hfiber f b <~> Empty

In O (hfiber f b)
exact (inO_equiv_inO Empty e^-1). Defined. (** Any map between modal types is modal. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O A
H0: In O B

MapIn O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: In O A
H0: In O B

MapIn O f
intros b; exact _. Defined. (** Modal maps cancel on the left. *)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C

MapIn O g -> MapIn O (g o f) -> MapIn O f
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C

MapIn O g -> MapIn O (g o f) -> MapIn O f
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
X: MapIn O g
X0: MapIn O (g o f)
b: B

In O (hfiber f b)
refine (inO_equiv_inO _ (hfiber_hfiber_compose_map f g b)). Defined. (** Modal maps also cancel with equivalences on the other side. *)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsEquiv f
H0: MapIn O (g o f)

MapIn O g
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsEquiv f
H0: MapIn O (g o f)

MapIn O g
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsEquiv f
H0: MapIn O (g o f)
b: C

In O (hfiber g b)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsEquiv f
H0: MapIn O (g o f)
b: C

hfiber (g o f) b <~> hfiber g b
exact (equiv_functor_sigma f (fun a => 1%equiv)). Defined. Definition cancelR_equiv_mapinO {A B C : Type} (f : A <~> B) (g : B -> C) `{MapIn O _ _ (g o f)} : MapIn O g := cancelR_isequiv_mapinO f g. (** The pullback of a modal map is modal. *)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: MapIn O g

MapIn O (f^* g)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: MapIn O g

MapIn O (f^* g)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: MapIn O g
b: B

In O (hfiber (f^* g) b)
refine (inO_equiv_inO _ (hfiber_pullback_along f g b)^-1). Defined.
O: ReflectiveSubuniverse
A, B, C: Type
g: C -> A
f: B -> A
H: MapIn O f

MapIn O ((g ^*') f)
O: ReflectiveSubuniverse
A, B, C: Type
g: C -> A
f: B -> A
H: MapIn O f

MapIn O ((g ^*') f)
O: ReflectiveSubuniverse
A, B, C: Type
g: C -> A
f: B -> A
H: MapIn O f
c: C

In O (hfiber ((g ^*') f) c)
refine (inO_equiv_inO _ (hfiber_pullback_along' g f c)^-1). Defined. (** [functor_sum] preserves modal maps. *)
O: ReflectiveSubuniverse
A, A', B, B': Type
f: A -> A'
g: B -> B'
H: MapIn O f
H0: MapIn O g

MapIn O (functor_sum f g)
O: ReflectiveSubuniverse
A, A', B, B': Type
f: A -> A'
g: B -> B'
H: MapIn O f
H0: MapIn O g

MapIn O (functor_sum f g)
O: ReflectiveSubuniverse
A, A', B, B': Type
f: A -> A'
g: B -> B'
H: MapIn O f
H0: MapIn O g
a: A'

In O (hfiber (functor_sum f g) (inl a))
O: ReflectiveSubuniverse
A, A', B, B': Type
f: A -> A'
g: B -> B'
H: MapIn O f
H0: MapIn O g
b: B'
In O (hfiber (functor_sum f g) (inr b))
O: ReflectiveSubuniverse
A, A', B, B': Type
f: A -> A'
g: B -> B'
H: MapIn O f
H0: MapIn O g
a: A'

In O (hfiber (functor_sum f g) (inl a))
refine (inO_equiv_inO _ (hfiber_functor_sum_l f g a)^-1).
O: ReflectiveSubuniverse
A, A', B, B': Type
f: A -> A'
g: B -> B'
H: MapIn O f
H0: MapIn O g
b: B'

In O (hfiber (functor_sum f g) (inr b))
refine (inO_equiv_inO _ (hfiber_functor_sum_r f g b)^-1). Defined. (** So does [unfunctor_sum], if both summands are preserved. These can't be [Instance]s since they require [Ha] and [Hb] to be supplied. *)
O: ReflectiveSubuniverse
A, A', B, B': Type
h: A + B -> A' + B'
Ha: forall a : A, is_inl (h (inl a))
Hb: forall b : B, is_inr (h (inr b))
H: MapIn O h

MapIn O (unfunctor_sum_l h Ha)
O: ReflectiveSubuniverse
A, A', B, B': Type
h: A + B -> A' + B'
Ha: forall a : A, is_inl (h (inl a))
Hb: forall b : B, is_inr (h (inr b))
H: MapIn O h

MapIn O (unfunctor_sum_l h Ha)
O: ReflectiveSubuniverse
A, A', B, B': Type
h: A + B -> A' + B'
Ha: forall a : A, is_inl (h (inl a))
Hb: forall b : B, is_inr (h (inr b))
H: MapIn O h
a: A'

In O (hfiber (unfunctor_sum_l h Ha) a)
refine (inO_equiv_inO _ (hfiber_unfunctor_sum_l h Ha Hb a)^-1). Defined.
O: ReflectiveSubuniverse
A, A', B, B': Type
h: A + B -> A' + B'
Ha: forall a : A, is_inl (h (inl a))
Hb: forall b : B, is_inr (h (inr b))
H: MapIn O h

MapIn O (unfunctor_sum_r h Hb)
O: ReflectiveSubuniverse
A, A', B, B': Type
h: A + B -> A' + B'
Ha: forall a : A, is_inl (h (inl a))
Hb: forall b : B, is_inr (h (inr b))
H: MapIn O h

MapIn O (unfunctor_sum_r h Hb)
O: ReflectiveSubuniverse
A, A', B, B': Type
h: A + B -> A' + B'
Ha: forall a : A, is_inl (h (inl a))
Hb: forall b : B, is_inr (h (inr b))
H: MapIn O h
b: B'

In O (hfiber (unfunctor_sum_r h Hb) b)
refine (inO_equiv_inO _ (hfiber_unfunctor_sum_r h Ha Hb b)^-1). Defined. End ModalMaps. (** ** Modally connected maps *) (** Connectedness of a map can again be defined in two equivalent ways: by connectedness of its fibers (as types), or by the lifting property/elimination principle against modal types. We use the former; the equivalence with the latter is given below in [conn_map_elim], [conn_map_comp], and [conn_map_from_extension_elim]. *) Class IsConnMap (O : ReflectiveSubuniverse@{i}) {A : Type@{i}} {B : Type@{i}} (f : A -> B) := isconnected_hfiber_conn_map (** The extra universe [k] is >= max(i,j). *) : forall b:B, IsConnected@{i} O (hfiber@{i i} f b). Global Existing Instance isconnected_hfiber_conn_map. Section ConnectedMaps. Universe i. Context (O : ReflectiveSubuniverse@{i}). (** Any equivalence is connected *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

IsConnMap O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsEquiv f

IsConnMap O f
intros b; exact _. Defined. (** Anything homotopic to a connected map is connected. *)
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
h: f == g

IsConnMap O f -> IsConnMap O g
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
h: f == g

IsConnMap O f -> IsConnMap O g
O: ReflectiveSubuniverse
A, B: Type
f, g: A -> B
h: f == g
X: IsConnMap O f
b: B

IsConnected O (hfiber g b)
exact (isconnected_equiv O (hfiber@{i i} f b) (equiv_hfiber_homotopic@{i i i} f g h b) _). Defined. (** The pullback of a connected map is connected *)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: IsConnMap O g

IsConnMap O (f^* g)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: IsConnMap O g

IsConnMap O (f^* g)
O: ReflectiveSubuniverse
A, B, C: Type
f: B -> A
g: C -> A
H: IsConnMap O g
b: B

IsConnected O (hfiber (f^* g) b)
refine (isconnected_equiv _ _ (hfiber_pullback_along f g b)^-1 _). Defined.
O: ReflectiveSubuniverse
A, B, C: Type
g: C -> A
f: B -> A
H: IsConnMap O f

IsConnMap O ((g ^*') f)
O: ReflectiveSubuniverse
A, B, C: Type
g: C -> A
f: B -> A
H: IsConnMap O f

IsConnMap O ((g ^*') f)
O: ReflectiveSubuniverse
A, B, C: Type
g: C -> A
f: B -> A
H: IsConnMap O f
c: C

IsConnected O (hfiber ((g ^*') f) c)
refine (isconnected_equiv _ _ (hfiber_pullback_along' g f c)^-1 _). Defined. (** The projection from a family of connected types is connected. *)
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: forall a : A, IsConnected O (B a)

IsConnMap O pr1
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: forall a : A, IsConnected O (B a)

IsConnMap O pr1
O: ReflectiveSubuniverse
A: Type
B: A -> Type
H: forall a : A, IsConnected O (B a)
a: A

IsConnected O (hfiber pr1 a)
refine (isconnected_equiv O (B a) (hfiber_fibration a B) _). Defined. (** Being connected is an hprop *)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B

IsHProp (IsConnMap O f)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B

IsHProp (IsConnMap O f)
apply istrunc_forall. Defined. (** Connected maps are orthogonal to modal maps (i.e. familes of modal types). *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)

forall b : B, P b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)

forall b : B, P b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)
b: B

P b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)
b: B

hfiber f b -> P b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)
b: B
a: A
p: f a = b

P b
exact (transport P p (d a)). Defined.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)

forall a : A, conn_map_elim f P d (f a) = d a
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)

forall a : A, conn_map_elim f P d (f a) = d a
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)
a: A

conn_map_elim f P d (f a) = d a
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)
a: A

(isconnected_elim O (P (f a)) (fun X : hfiber f (f a) => transport P X.2 (d X.1))).1 = d a
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)
a: A
fibermap:= fun a0p : hfiber f (f a) => let a0 := a0p.1 in let p := a0p.2 in transport P p (d a0): hfiber f (f a) -> P (f a)

(isconnected_elim O (P (f a)) fibermap).1 = d a
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)
a: A
fibermap:= fun a0p : hfiber f (f a) => let a0 := a0p.1 in let p := a0p.2 in transport P p (d a0): hfiber f (f a) -> P (f a)
x: P (f a)
e: forall x0 : hfiber f (f a), fibermap x0 = x

(x; e).1 = d a
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)
a: A
fibermap:= fun a0p : hfiber f (f a) => let a0 := a0p.1 in let p := a0p.2 in transport P p (d a0): hfiber f (f a) -> P (f a)
x: P (f a)
e: forall x0 : hfiber f (f a), fibermap x0 = x

(x; e).1 = fibermap (a; 1)
apply inverse, e. Defined. (** A map which is both connected and modal is an equivalence. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
H0: MapIn O f

IsEquiv f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
H0: MapIn O f

IsEquiv f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
H0: MapIn O f

IsTruncMap (-2) f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
H0: MapIn O f
b: B

Contr (hfiber f b)
apply (contr_trunc_conn O). Defined. (** We can re-express this in terms of extensions. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)

ExtensionAlong f P d
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)

ExtensionAlong f P d
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)
d: forall a : A, P (f a)

forall x : A, conn_map_elim f P d (f x) = d x
apply conn_map_comp. Defined.
O: ReflectiveSubuniverse
n: nat
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)

ExtendableAlong n f P
O: ReflectiveSubuniverse
n: nat
A, B: Type
f: A -> B
H: IsConnMap O f
P: B -> Type
H0: forall b : B, In O (P b)

ExtendableAlong n f P
O: ReflectiveSubuniverse
n: nat
A, B: Type
f: A -> B
H: IsConnMap O f

forall P : B -> Type, (forall b : B, In O (P b)) -> ExtendableAlong n f P
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
n: nat
IHn: forall P : B -> Type, (forall b : B, In O (P b)) -> ExtendableAlong n f P
P: B -> Type
H0: forall b : B, In O (P b)

forall g : forall a : A, P (f a), ExtensionAlong f P g
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
n: nat
IHn: forall P : B -> Type, (forall b : B, In O (P b)) -> ExtendableAlong n f P
P: B -> Type
H0: forall b : B, In O (P b)
forall h k : forall b : B, P b, ExtendableAlong n f (fun b : B => h b = k b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
n: nat
IHn: forall P : B -> Type, (forall b : B, In O (P b)) -> ExtendableAlong n f P
P: B -> Type
H0: forall b : B, In O (P b)

forall g : forall a : A, P (f a), ExtensionAlong f P g
intros d; apply extension_conn_map_elim; exact _.
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f
n: nat
IHn: forall P : B -> Type, (forall b : B, In O (P b)) -> ExtendableAlong n f P
P: B -> Type
H0: forall b : B, In O (P b)

forall h k : forall b : B, P b, ExtendableAlong n f (fun b : B => h b = k b)
intros h k; apply IHn; exact _. Defined. Definition ooextendable_conn_map_inO {A B : Type} (f : A -> B) `{IsConnMap O _ _ f} (P : B -> Type) `{forall b:B, In O (P b)} : ooExtendableAlong f P := fun n => extendable_conn_map_inO n f P.
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall a : A, P (f a)
e, e': ExtensionAlong f P d

e = e'
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall a : A, P (f a)
e, e': ExtensionAlong f P d

e = e'
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall a : A, P (f a)
e, e': ExtensionAlong f P d

ExtensionAlong f (fun y : B => e.1 y = e'.1 y) (fun x : A => e.2 x @ (e'.2 x)^)
refine (extension_conn_map_elim _ _ _). Defined. (** It follows that [conn_map_elim] is actually an equivalence. *)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)

IsEquiv (fun g : forall b : B, P b => g oD f)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)

IsEquiv (fun g : forall b : B, P b => g oD f)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

Contr (hfiber (fun g : forall b : B, P b => g oD f) d)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

IsHProp (hfiber (fun g : forall b : B, P b => g oD f) d)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)
hfiber (fun g : forall b : B, P b => g oD f) d
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

IsHProp (hfiber (fun g : forall b : B, P b => g oD f) d)
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

{g : forall b : B, P b & g oD f == d} <~> hfiber (fun g : forall b : B, P b => g oD f) d
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)
IsHProp {g : forall b : B, P b & g oD f == d}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

{g : forall b : B, P b & g oD f == d} <~> hfiber (fun g : forall b : B, P b => g oD f) d
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)
g: forall b : B, P b

g oD f == d <~> g oD f = d
apply equiv_path_forall.
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

IsHProp {g : forall b : B, P b & g oD f == d}
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

forall x y : {g : forall b : B, P b & g oD f == d}, x = y
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)
g, h: {g : forall b : B, P b & g oD f == d}

g = h
exact (allpath_extension_conn_map f P d g h).
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

hfiber (fun g : forall b : B, P b => g oD f) d
O: ReflectiveSubuniverse
H: Funext
A, B: Type
f: A -> B
H0: IsConnMap O f
P: B -> Type
H1: forall b : B, In O (P b)
d: forall x : A, P (f x)

conn_map_elim f P d oD f = d
apply path_forall; intros x; apply conn_map_comp. Defined. Definition equiv_o_conn_map `{Funext} {A B : Type} (f : A -> B) `{IsConnMap O _ _ f} (P : B -> Type) `{forall b:B, In O (P b)} : (forall b, P b) <~> (forall a, P (f a)) := Build_Equiv _ _ _ (isequiv_o_conn_map f P). (** When considering lexness properties, we often want to consider the property of the universe of modal types being modal. We can't say this directly (except in the accessible, hence liftable, case) because it lives in a higher universe, but we can make a direct extendability statement. Here we prove a lemma that oo-extendability into the universe follows from plain extendability, essentially because the type of equivalences between two [O]-modal types is [O]-modal. *)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P

ooExtendableAlong f (fun _ : B => Type_ O)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P

ooExtendableAlong f (fun _ : B => Type_ O)
(** By definition, in addition to our assumption [extP] that maps into [Type_ O] extend along [f], we must show that sections of families of equivalences are [ooExtendableAlong] it. *)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P

ExtendableAlong 0 f (fun _ : B => Type_ O)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P
ExtendableAlong 1 f (fun _ : B => Type_ O)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P
n: nat
ExtendableAlong n.+2 f (fun _ : B => Type_ O)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P

ExtendableAlong 0 f (fun _ : B => Type_ O)
exact tt. (* n = 0 *) (** Note that due to the implementation of [ooExtendableAlong], we actually have to use [extP] twice (there should probably be a general cofixpoint lemma for this). *)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P

ExtendableAlong 1 f (fun _ : B => Type_ O)
split; [ apply extP | intros; exact tt ]. (* n = 1 *)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P
n: nat

ExtendableAlong n.+2 f (fun _ : B => Type_ O)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P
n: nat

forall h k : B -> Type_ O, ExtendableAlong n.+1 f (fun b : B => h b = k b)
(** What remains is to extend families of paths. *)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P
n: nat
P, Q: B -> Type_ O

forall b : B, (P b <~> Q b) <~> P b = Q b
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P
n: nat
P, Q: B -> Type_ O
ooExtendableAlong f (fun b : B => P b <~> Q b)
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P
n: nat
P, Q: B -> Type_ O

forall b : B, (P b <~> Q b) <~> P b = Q b
intros x; refine (equiv_path_TypeO _ _ _ oE equiv_path_universe _ _).
O: ReflectiveSubuniverse
H: Univalence
A, B: Type
f: A -> B
H0: IsConnMap O f
extP: forall P : A -> Type_ O, ExtensionAlong f (fun _ : B => Type_ O) P
n: nat
P, Q: B -> Type_ O

ooExtendableAlong f (fun b : B => P b <~> Q b)
rapply ooextendable_conn_map_inO. Defined. (** Conversely, if a map satisfies this elimination principle (expressed via extensions), then it is connected. This completes the proof of Lemma 7.5.7 from the book. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B

(forall P : B -> Type, (forall b : B, In O (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d) -> IsConnMap O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B

(forall P : B -> Type, (forall b : B, In O (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d) -> IsConnMap O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
Hf: forall P : B -> Type, (forall b : B, In O (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d
b: B

IsConnected O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
Hf: forall P : B -> Type, (forall b : B, In O (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d
b: B

NullHomotopy (to O (hfiber f b))
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
Hf: forall P : B -> Type, (forall b : B, In O (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d
b: B
e: ExtensionAlong f (fun b : B => O (hfiber f b)) (fun a : A => to O (hfiber f (f a)) (a; 1))

NullHomotopy (to O (hfiber f b))
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
Hf: forall P : B -> Type, (forall b : B, In O (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d
b: B
e: ExtensionAlong f (fun b : B => O (hfiber f b)) (fun a : A => to O (hfiber f (f a)) (a; 1))

forall x : hfiber f b, to O (hfiber f b) x = e.1 b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
Hf: forall P : B -> Type, (forall b : B, In O (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d
b: B
e: ExtensionAlong f (fun b : B => O (hfiber f b)) (fun a : A => to O (hfiber f (f a)) (a; 1))
a: A
p: f a = b

to O (hfiber f b) (a; p) = e.1 b
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
Hf: forall P : B -> Type, (forall b : B, In O (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d
e: ExtensionAlong f (fun b : B => O (hfiber f b)) (fun a : A => to O (hfiber f (f a)) (a; 1))
a: A

to O (hfiber f (f a)) (a; 1) = e.1 (f a)
symmetry; apply (e.2). Defined. (** Lemma 7.5.6: Connected maps compose and cancel on the right. *)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O g

IsConnMap O (g o f)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O g

IsConnMap O (g o f)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O g
P: C -> Type
P_inO: forall b : C, In O (P b)
d: forall a : A, P (g (f a))

ExtensionAlong (fun x : A => g (f x)) P d
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O g
P: C -> Type
P_inO: forall b : C, In O (P b)
d: forall a : A, P (g (f a))
a: A

conn_map_elim g P (conn_map_elim f (fun b : B => P (g b)) d) (g (f a)) = d a
exact (conn_map_comp g P _ _ @ conn_map_comp f (fun b => P (g b)) d a). Defined.
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O (g o f)

IsConnMap O g
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O (g o f)

IsConnMap O g
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O (g o f)
P: C -> Type
P_inO: forall b : C, In O (P b)
d: forall a : B, P (g a)

ExtensionAlong g P d
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O (g o f)
P: C -> Type
P_inO: forall b : C, In O (P b)
d: forall a : B, P (g a)
b: B

conn_map_elim (fun x : A => g (f x)) P (d oD f) (g b) = d b
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsConnMap O f
H0: IsConnMap O (g o f)
P: C -> Type
P_inO: forall b : C, In O (P b)
d: forall a : B, P (g a)
b: B
a: A

conn_map_elim (fun x : A => g (f x)) P (d oD f) (g (f a)) = d (f a)
exact (conn_map_comp (g o f) P (d oD f) a). Defined. (** Connected maps also cancel with equivalences on the other side. *)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsEquiv g
H0: IsConnMap O (g o f)

IsConnMap O f
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsEquiv g
H0: IsConnMap O (g o f)

IsConnMap O f
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsEquiv g
H0: IsConnMap O (g o f)
b: B

IsConnected O (hfiber f b)
O: ReflectiveSubuniverse
A, B, C: Type
f: A -> B
g: B -> C
H: IsEquiv g
H0: IsConnMap O (g o f)
b: B

hfiber (g o f) (g b) <~> hfiber f b
exact (equiv_inverse (equiv_functor_sigma_id (fun a => equiv_ap g (f a) b))). Defined. Definition cancelL_equiv_conn_map {A B C : Type} (f : A -> B) (g : B <~> C) `{IsConnMap O _ _ (g o f)} : IsConnMap O f := cancelL_isequiv_conn_map f g. (** The constant map to [Unit] is connected just when its domain is. *)
O: ReflectiveSubuniverse
A: Type
H: IsConnMap O (const_tt A)

IsConnected O A
O: ReflectiveSubuniverse
A: Type
H: IsConnMap O (const_tt A)

IsConnected O A
refine (isconnected_equiv O (hfiber (const_tt A) tt) (equiv_sigma_contr _) _). Defined. #[local] Hint Immediate isconnected_conn_map_to_unit : typeclass_instances.
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A

IsConnMap O (const_tt A)
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A

IsConnMap O (const_tt A)
O: ReflectiveSubuniverse
A: Type
H: IsConnected O A
u: Unit

IsConnected O (hfiber (const_tt A) u)
refine (isconnected_equiv O A (equiv_sigma_contr _)^-1 _). Defined. (* Lemma 7.5.10: A map to a type in [O] exhibits its codomain as the [O]-reflection of its domain if it is [O]-connected. (The converse is true if and only if [O] is a modality.) *)
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f

IsEquiv (O_rec f)
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f

IsEquiv (O_rec f)
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f

(fun x : B => O_rec f (conn_map_elim f (fun _ : B => O A) (to O A) x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f
(fun x : O A => conn_map_elim f (fun _ : B => O A) (to O A) (O_rec f x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f

(fun x : B => O_rec f (conn_map_elim f (fun _ : B => O A) (to O A) x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f
x: B

O_rec f (conn_map_elim f (fun _ : B => O A) (to O A) x) = x
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f
x: B

(fun b : B => O_rec f (conn_map_elim f (fun _ : B => O A) (to O A) b) = b) x
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f
x: B
a: A

O_rec f (conn_map_elim f (fun _ : B => O A) (to O A) (f a)) = f a
exact (ap (O_rec f) (conn_map_comp f (fun _ => O A) (to O A) a) @ O_rec_beta f a).
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f

(fun x : O A => conn_map_elim f (fun _ : B => O A) (to O A) (O_rec f x)) == idmap
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f
a: A

conn_map_elim f (fun _ : B => O A) (to O A) (O_rec f (to O A a)) = to O A a
O: ReflectiveSubuniverse
A, B: Type
H: In O B
f: A -> B
H0: IsConnMap O f
a: A

conn_map_elim f (fun _ : B => O A) (to O A) (f a) = to O A a
refine (conn_map_comp f (fun _ => O A) (to O A) a). Defined. (** Lemma 7.5.12 *) Section ConnMapFunctorSigma. Context {A B : Type} {P : A -> Type} {Q : B -> Type} (f : A -> B) (g : forall a, P a -> Q (f a)) `{forall a, IsConnMap O (g a)}.
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

O (hfiber (functor_sigma f g) (b; v)) <~> O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

O (hfiber (functor_sigma f g) (b; v)) <~> O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

O (hfiber (functor_sigma f g) (b; v)) <~> O {w : hfiber f b & hfiber (g w.1) (transport Q (w.2)^ v)}
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b
O {w : hfiber f b & hfiber (g w.1) (transport Q (w.2)^ v)} <~> O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

O (hfiber (functor_sigma f g) (b; v)) <~> O {w : hfiber f b & hfiber (g w.1) (transport Q (w.2)^ v)}
apply equiv_O_functor, hfiber_functor_sigma.
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

O {w : hfiber f b & hfiber (g w.1) (transport Q (w.2)^ v)} <~> O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

O {w : hfiber f b & hfiber (g w.1) (transport Q (w.2)^ v)} <~> O {w : hfiber f b & O (hfiber (g w.1) (transport Q (w.2)^ v))}
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b
O {w : hfiber f b & O (hfiber (g w.1) (transport Q (w.2)^ v))} <~> O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

O {w : hfiber f b & hfiber (g w.1) (transport Q (w.2)^ v)} <~> O {w : hfiber f b & O (hfiber (g w.1) (transport Q (w.2)^ v))}
symmetry; apply equiv_O_sigma_O.
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

O {w : hfiber f b & O (hfiber (g w.1) (transport Q (w.2)^ v))} <~> O (hfiber f b)
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
b: B
v: Q b

{w : hfiber f b & O (hfiber (g w.1) (transport Q (w.2)^ v))} <~> hfiber f b
apply equiv_sigma_contr; intros [a p]; simpl; exact _. Defined.
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
H0: IsConnMap O f

IsConnMap O (functor_sigma f g)
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
H0: IsConnMap O f

IsConnMap O (functor_sigma f g)
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
H0: IsConnMap O f
b: B
v: Q b

IsConnected O (hfiber (functor_sigma f g) (b; v))
refine (contr_equiv' _ (equiv_inverse (equiv_O_hfiber_functor_sigma b v))). Defined.
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
inh: forall b : B, Q b
H0: IsConnMap O (functor_sigma f g)

IsConnMap O f
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
inh: forall b : B, Q b
H0: IsConnMap O (functor_sigma f g)

IsConnMap O f
O: ReflectiveSubuniverse
A, B: Type
P: A -> Type
Q: B -> Type
f: A -> B
g: forall a : A, P a -> Q (f a)
H: forall a : A, IsConnMap O (g a)
inh: forall b : B, Q b
H0: IsConnMap O (functor_sigma f g)
b: B

IsConnected O (hfiber f b)
refine (contr_equiv _ (equiv_O_hfiber_functor_sigma b (inh b))). Defined. End ConnMapFunctorSigma. (** Lemma 7.5.13. The "if" direction is a special case of [conn_map_functor_sigma], so we prove only the "only if" direction. *)
O: ReflectiveSubuniverse
A: Type
P, Q: A -> Type
f: forall a : A, P a -> Q a
H: IsConnMap O (functor_sigma idmap f)

forall a : A, IsConnMap O (f a)
O: ReflectiveSubuniverse
A: Type
P, Q: A -> Type
f: forall a : A, P a -> Q a
H: IsConnMap O (functor_sigma idmap f)

forall a : A, IsConnMap O (f a)
O: ReflectiveSubuniverse
A: Type
P, Q: A -> Type
f: forall a : A, P a -> Q a
H: IsConnMap O (functor_sigma idmap f)
a: A
q: Q a

IsConnected O (hfiber (f a) q)
O: ReflectiveSubuniverse
A: Type
P, Q: A -> Type
f: forall a : A, P a -> Q a
H: IsConnMap O (functor_sigma idmap f)
a: A
q: Q a

hfiber (functor_sigma idmap f) (a; q) <~> hfiber (f a) q
exact (hfiber_functor_sigma_idmap P Q f a q). Defined. (** Lemma 7.5.14: Connected maps are inverted by [O]. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f

O_inverts O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f

O_inverts O f
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f

forall Z : Type, In O Z -> ExtendableAlong 2 f (fun _ : B => Z)
intros; rapply extendable_conn_map_inO. Defined. (** As a consequence, connected maps between modal types are equivalences. *) Definition isequiv_conn_map_ino {A B : Type} (f : A -> B) `{In O A} `{In O B} `{IsConnMap O _ _ f} : IsEquiv f := isequiv_commsq' f (O_functor O f) (to O A) (to O B) (to_O_natural O f). (** Connectedness is preserved by [O_functor]. *)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f

IsConnMap O (O_functor O f)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f

IsConnMap O (O_functor O f)
O: ReflectiveSubuniverse
A, B: Type
f: A -> B
H: IsConnMap O f

IsConnMap O (O_rec (fun x : A => to O B (f x)))
rapply conn_map_compose. Defined. (** Connected maps are preserved by coproducts *)
O: ReflectiveSubuniverse
A, B, A', B': Type
f: A -> A'
g: B -> B'
H: IsConnMap O f
H0: IsConnMap O g

IsConnMap O (functor_sum f g)
O: ReflectiveSubuniverse
A, B, A', B': Type
f: A -> A'
g: B -> B'
H: IsConnMap O f
H0: IsConnMap O g

IsConnMap O (functor_sum f g)
O: ReflectiveSubuniverse
A, B, A', B': Type
f: A -> A'
g: B -> B'
H: IsConnMap O f
H0: IsConnMap O g
P: A' + B' -> Type
P_inO: forall b : A' + B', In O (P b)
d: forall a : A + B, P (functor_sum f g a)

ExtensionAlong (functor_sum f g) P d
apply extension_functor_sum; rapply ooextendable_conn_map_inO. Defined. (** Connected maps are preserved by coequalizers *)
O: ReflectiveSubuniverse
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H: IsConnMap O k
H0: IsConnMap O h

IsConnMap O (functor_coeq h k p q)
O: ReflectiveSubuniverse
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H: IsConnMap O k
H0: IsConnMap O h

IsConnMap O (functor_coeq h k p q)
O: ReflectiveSubuniverse
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H: IsConnMap O k
H0: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)

ExtensionAlong (functor_coeq h k p q) P d
O: ReflectiveSubuniverse
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H: IsConnMap O k
H0: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)

ExtendableAlong 1 k (fun x : A' => P (coeq x))
O: ReflectiveSubuniverse
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H: IsConnMap O k
H0: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)
forall u v : forall x : B', P (coeq (g' x)), ExtendableAlong 1 h (fun x : B' => u x = v x)
O: ReflectiveSubuniverse
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H: IsConnMap O k
H0: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)

ExtendableAlong 1 k (fun x : A' => P (coeq x))
rapply ooextendable_conn_map_inO.
O: ReflectiveSubuniverse
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H: IsConnMap O k
H0: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)

forall u v : forall x : B', P (coeq (g' x)), ExtendableAlong 1 h (fun x : B' => u x = v x)
intros; rapply ooextendable_conn_map_inO. Defined. (** And by pushouts *)
O: ReflectiveSubuniverse
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H: IsConnMap O h
H0: IsConnMap O k
H1: IsConnMap O l

IsConnMap O (functor_pushout h k l p q)
O: ReflectiveSubuniverse
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H: IsConnMap O h
H0: IsConnMap O k
H1: IsConnMap O l

IsConnMap O (functor_pushout h k l p q)
O: ReflectiveSubuniverse
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H: IsConnMap O h
H0: IsConnMap O k
H1: IsConnMap O l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)

ExtensionAlong (functor_pushout h k l p q) P d
O: ReflectiveSubuniverse
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H: IsConnMap O h
H0: IsConnMap O k
H1: IsConnMap O l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)

ExtendableAlong 1 (functor_sum k l) (fun x : B' + C' => P (coeq x))
O: ReflectiveSubuniverse
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H: IsConnMap O h
H0: IsConnMap O k
H1: IsConnMap O l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)
forall u v : forall x : A', P (coeq (inr (g' x))), ExtendableAlong 1 h (fun x : A' => u x = v x)
O: ReflectiveSubuniverse
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H: IsConnMap O h
H0: IsConnMap O k
H1: IsConnMap O l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)

ExtendableAlong 1 (functor_sum k l) (fun x : B' + C' => P (coeq x))
apply extendable_functor_sum; rapply ooextendable_conn_map_inO.
O: ReflectiveSubuniverse
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H: IsConnMap O h
H0: IsConnMap O k
H1: IsConnMap O l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)

forall u v : forall x : A', P (coeq (inr (g' x))), ExtendableAlong 1 h (fun x : A' => u x = v x)
intros; rapply ooextendable_conn_map_inO. Defined. End ConnectedMaps. (** ** Containment of (reflective) subuniverses *) (** One subuniverse is contained in another if every [O1]-modal type is [O2]-modal. We define this parametrized by three universes: [O1] and [O2] are reflective subuniverses of [Type@{i1}] and [Type@{i2}] respectively, and the relation says that all types in [Type@{j}] that [O1]-modal are also [O2]-modal. This implies [j <= i1] and [j <= i2], of course. The most common application is when [i1 = i2 = j], but it's sometimes useful to talk about a subuniverse of a larger universe agreeing with a subuniverse of a smaller universe on the smaller universe. *) Class O_leq@{i1 i2 j} (O1 : Subuniverse@{i1}) (O2 : Subuniverse@{i2}) := inO_leq : forall (A : Type@{j}), In O1 A -> In O2 A. Arguments inO_leq O1 O2 {_} A _. Declare Scope subuniverse_scope. Notation "O1 <= O2" := (O_leq O1 O2) : subuniverse_scope. Open Scope subuniverse_scope.

Reflexive O_leq

Reflexive O_leq
intros O A ?; assumption. Defined.

Transitive O_leq

Transitive O_leq
O1, O2, O3: Subuniverse
O12: O1 <= O2
O23: O2 <= O3
A: Type
X: In O1 A

In O3 A
O1, O2, O3: Subuniverse
O12: O1 <= O2
O23: O2 <= O3
A: Type
X: In O1 A

In O2 A
rapply (@inO_leq O1 O2). Defined.
O1, O2: Subuniverse
H: O1 <= O2
A, B: Type
f: A -> B
H0: MapIn O1 f

MapIn O2 f
O1, O2: Subuniverse
H: O1 <= O2
A, B: Type
f: A -> B
H0: MapIn O1 f

MapIn O2 f
intros b; rapply (inO_leq O1 O2). Defined. (** This implies that every [O2]-connected type is [O1]-connected, and similarly for maps and equivalences. We give universe annotations so that [O1] and [O2] don't have to be on the same universe, but we do have to have [i1 <= i2] for this statement. When [i2 <= i1] it seems that the statement might not be true unless the RSU on the larger universe is accessibly extended from the smaller one; see [Localization.v]. *)
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A: Type
H0: IsConnected O2 A

IsConnected O1 A
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A: Type
H0: IsConnected O2 A

IsConnected O1 A
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A: Type
H0: IsConnected O2 A

forall C : Type, In O1 C -> forall f : A -> C, NullHomotopy f
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A: Type
H0: IsConnected O2 A
C: Type
C1: In O1 C
f: A -> C

NullHomotopy f
apply (isconnected_elim O2); srapply inO_leq; exact _. Defined. (** This one has the same universe constraint [i1 <= i2]. *)
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
H0: IsConnMap O2 f

IsConnMap O1 f
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
H0: IsConnMap O2 f

IsConnMap O1 f
(** We could prove this by applying [isconnected_O_leq] fiberwise, but unless we were very careful that would collapse the two universes [i1] and [i2]. So instead we just give an analogous direct proof. *)
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
H0: IsConnMap O2 f

forall P : B -> Type, (forall b : B, In O1 (P b)) -> forall d : forall a : A, P (f a), ExtensionAlong f P d
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
H0: IsConnMap O2 f
P: B -> Type
P_inO: forall b : B, In O1 (P b)
g: forall a : A, P (f a)

ExtensionAlong f P g
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
H0: IsConnMap O2 f
P: B -> Type
P_inO: forall b : B, In O1 (P b)
g: forall a : A, P (f a)

forall b : B, In O2 (P b)
intros b; rapply (inO_leq O1). Defined. (** This is Lemma 2.12(i) in CORS, again with the same universe constraint [i1 <= i2]. *)
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
O_inverts0: O_inverts O2 f

O_inverts O1 f
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
O_inverts0: O_inverts O2 f

O_inverts O1 f
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
O_inverts0: O_inverts O2 f
Z: Type
Z_inO: In O1 Z

ExtendableAlong 2 f (fun _ : B => Z)
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
O_inverts0: O_inverts O2 f
Z: Type
Z_inO: In O1 Z
i:= inO_leq O1 O2 Z Z_inO: In O2 Z

ExtendableAlong 2 f (fun _ : B => Z)
O1, O2: ReflectiveSubuniverse
H: O1 <= O2
A, B: Type
f: A -> B
O_inverts0: O_inverts O2 f
Z: Type
Z_inO: In O1 Z
i:= inO_leq O1 O2 Z Z_inO: In O2 Z

ExtendableAlong 2 f (fun _ : B => Z)
apply (ooextendable_O_inverts O2); exact _. Defined. (** ** Equality of (reflective) subuniverses *) (** Two subuniverses are the same if they have the same modal types. The universe parameters are the same as for [O_leq]: [O1] and [O2] are reflective subuniverses of [Type@{i1}] and [Type@{i2}], and the relation says that they agree when restricted to [Type@{j}], where [j <= i1] and [j <= i2]. *) Class O_eq@{i1 i2 j} (O1 : Subuniverse@{i1}) (O2 : Subuniverse@{i2}) := { O_eq_l : O_leq@{i1 i2 j} O1 O2 ; O_eq_r : O_leq@{i2 i1 j} O2 O1 ; }. Global Existing Instances O_eq_l O_eq_r. Infix "<=>" := O_eq : subuniverse_scope. Definition issig_O_eq O1 O2 : _ <~> O_eq O1 O2 := ltac:(issig).

Reflexive O_eq

Reflexive O_eq
intros; split; reflexivity. Defined.

Transitive O_eq

Transitive O_eq
intros O1 O2 O3; split; refine (transitivity (y := O2) _ _). Defined.

Symmetric O_eq

Symmetric O_eq
intros O1 O2 [? ?]; split; assumption. Defined. Definition issig_subuniverse : _ <~> Subuniverse := ltac:(issig).
H: Univalence
O1, O2: Subuniverse

O1 <=> O2 <~> O1 = O2
H: Univalence
O1, O2: Subuniverse

O1 <=> O2 <~> O1 = O2
H: Univalence
O1, O2: Subuniverse

{_ : O1 <= O2 & O2 <= O1} <~> O1 = O2
H: Univalence

forall b0 b1 : {H : Type -> Type & {_ : Funext -> forall T : Type, IsHProp (H T) & forall T U : Type, H T -> forall f : T -> U, IsEquiv f -> H U}}, {_ : issig_subuniverse b0 <= issig_subuniverse b1 & issig_subuniverse b1 <= issig_subuniverse b0} <~> b0 = b1
H: Univalence
O1, O2: {H : Type -> Type & {_ : Funext -> forall T : Type, IsHProp (H T) & forall T U : Type, H T -> forall f : T -> U, IsEquiv f -> H U}}

{_ : {| In_internal := O1.1; hprop_inO_internal := (O1.2).1; inO_equiv_inO_internal := (O1.2).2 |} <= {| In_internal := O2.1; hprop_inO_internal := (O2.2).1; inO_equiv_inO_internal := (O2.2).2 |} & {| In_internal := O2.1; hprop_inO_internal := (O2.2).1; inO_equiv_inO_internal := (O2.2).2 |} <= {| In_internal := O1.1; hprop_inO_internal := (O1.2).1; inO_equiv_inO_internal := (O1.2).2 |}} <~> O1 = O2
H: Univalence
O1, O2: {H : Type -> Type & {_ : Funext -> forall T : Type, IsHProp (H T) & forall T U : Type, H T -> forall f : T -> U, IsEquiv f -> H U}}

{_ : {| In_internal := O1.1; hprop_inO_internal := (O1.2).1; inO_equiv_inO_internal := (O1.2).2 |} <= {| In_internal := O2.1; hprop_inO_internal := (O2.2).1; inO_equiv_inO_internal := (O2.2).2 |} & {| In_internal := O2.1; hprop_inO_internal := (O2.2).1; inO_equiv_inO_internal := (O2.2).2 |} <= {| In_internal := O1.1; hprop_inO_internal := (O1.2).1; inO_equiv_inO_internal := (O1.2).2 |}} <~> O1.1 = O2.1
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U

{_ : {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |} <= {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} & {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} <= {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |}} <~> O1 = O2
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U

{_ : {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |} <= {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} & {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} <= {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |}} <~> O1 == O2
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U

IsHProp {_ : {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |} <= {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} & {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} <= {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |}}
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
{_ : {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |} <= {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} & {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} <= {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |}} -> O1 == O2
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
O1 == O2 -> {_ : {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |} <= {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} & {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} <= {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |}}
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U

IsHProp {_ : {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |} <= {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} & {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} <= {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |}}
srapply istrunc_sigma; unfold O_leq; exact _.
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U

{_ : {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |} <= {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} & {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} <= {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |}} -> O1 == O2
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
A: Type
h: O1 A -> O2 A
k: O2 A -> O1 A

O1 A = O2 A
apply path_universe_uncurried, equiv_iff_hprop; assumption.
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U

O1 == O2 -> {_ : {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |} <= {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} & {| In_internal := O2; hprop_inO_internal := O2h; inO_equiv_inO_internal := proj0 |} <= {| In_internal := O1; hprop_inO_internal := O1h; inO_equiv_inO_internal := proj2 |}}
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
A: Type
h: O1 A = O2 A
e: O1 A

O2 A
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
A: Type
h: O1 A = O2 A
e: O2 A
O1 A
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
A: Type
h: O1 A = O2 A
e: O1 A

O1 A
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
A: Type
h: O1 A = O2 A
e: O2 A
O1 A
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
A: Type
h: O1 A = O2 A
e: O1 A

O1 A
H: Univalence
O1: Type -> Type
O1h: Funext -> forall T : Type, IsHProp (O1 T)
proj2: forall T U : Type, O1 T -> forall f : T -> U, IsEquiv f -> O1 U
O2: Type -> Type
O2h: Funext -> forall T : Type, IsHProp (O2 T)
proj0: forall T U : Type, O2 T -> forall f : T -> U, IsEquiv f -> O2 U
A: Type
h: O1 A = O2 A
e: O2 A
O2 A
all:exact e. Defined. (** It should also be true that if [O1] and [O2] are reflective subuniverses, then [O1 <=> O2] is equivalent to [O1 = O2 :> ReflectiveSubuniverse]. Probably [contr_typeof_O_unit] should be useful in proving that. *) (** Reflections into one subuniverse are also reflections into an equal one. Unfortunately these almost certainly can't be [Instance]s for fear of infinite loops, since [<=>] is reflexive. *)
O1, O2: Subuniverse
H: O1 <= O2
A: Type
H0: PreReflects O1 A

PreReflects O2 A
O1, O2: Subuniverse
H: O1 <= O2
A: Type
H0: PreReflects O1 A

PreReflects O2 A
O1, O2: Subuniverse
H: O1 <= O2
A: Type
H0: PreReflects O1 A

Type
O1, O2: Subuniverse
H: O1 <= O2
A: Type
H0: PreReflects O1 A
In O2 ?O_reflector
O1, O2: Subuniverse
H: O1 <= O2
A: Type
H0: PreReflects O1 A
A -> ?O_reflector
O1, O2: Subuniverse
H: O1 <= O2
A: Type
H0: PreReflects O1 A

Type
exact (O_reflector O1 A).
O1, O2: Subuniverse
H: O1 <= O2
A: Type
H0: PreReflects O1 A

In O2 (O_reflector O1 A)
rapply (inO_leq O1 O2).
O1, O2: Subuniverse
H: O1 <= O2
A: Type
H0: PreReflects O1 A

A -> O_reflector O1 A
exact (to O1 A). Defined.
O1, O2: Subuniverse
H: O1 <=> O2
A: Type
H0: PreReflects O1 A
H1: Reflects O1 A

Reflects O2 A
O1, O2: Subuniverse
H: O1 <=> O2
A: Type
H0: PreReflects O1 A
H1: Reflects O1 A

Reflects O2 A
O1, O2: Subuniverse
H: O1 <=> O2
A: Type
H0: PreReflects O1 A
H1: Reflects O1 A
B: Type
B_inO2: In O2 B

ooExtendableAlong (to O2 A) (fun _ : O_reflector O2 A => B)
O1, O2: Subuniverse
H: O1 <=> O2
A: Type
H0: PreReflects O1 A
H1: Reflects O1 A
B: Type
B_inO2: In O2 B
X: In O1 B

ooExtendableAlong (to O2 A) (fun _ : O_reflector O2 A => B)
apply (extendable_to_O O1). Defined. (** ** Separated subuniverses *) (** For any subuniverse [O], a type is [O]-separated iff all its identity types are [O]-modal. We will study these further in [Separated.v], but we put the definition here because it's needed in [Descent.v]. *)
O: Subuniverse

Subuniverse
O: Subuniverse

Subuniverse
O: Subuniverse

Type -> Type
O: Subuniverse
Funext -> forall T : Type, IsHProp (?In_internal T)
O: Subuniverse
forall T U : Type, ?In_internal T -> forall f : T -> U, IsEquiv f -> ?In_internal U
O: Subuniverse

Type -> Type
intros A; exact (forall (x y:A), In O (x = y)).
O: Subuniverse

Funext -> forall T : Type, IsHProp ((fun A : Type => forall x y : A, In O (x = y)) T)
exact _.
O: Subuniverse

forall T U : Type, (fun A : Type => forall x y : A, In O (x = y)) T -> forall f : T -> U, IsEquiv f -> (fun A : Type => forall x y : A, In O (x = y)) U
O: Subuniverse
T, U: Type
T_inO: forall x y : T, In O (x = y)
f: T -> U
feq: IsEquiv f
x, y: U

In O (x = y)
refine (inO_equiv_inO' _ (equiv_ap f^-1 x y)^-1). Defined. Global Instance inO_paths_SepO (O : Subuniverse) {A : Type} {A_inO : In (Sep O) A} (x y : A) : In O (x = y) := A_inO x y. (** TODO: Where to put this? Morally it goes with the study of [<<] in [Modality.v] and [<<<] in [Descent.v] and [Sep] in [Separated.v], but it doesn't actually need any of those relations, only [O' <= Sep O], and it would also be nice to have it next to [O_inverts_functor_coeq]. It's a variation on the latter: if [O' <= Sep O], then for [O'] to invert [functor_coeq h k] it suffices that it invert [k] and that [h] be [O]-connected (by [conn_map_OO_inverts], which has different hypotheses but applies in many of the same examples, that is a weaker assumption). *)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h

O_inverts O' (functor_coeq h k p q)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h

O_inverts O' (functor_coeq h k p q)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h

forall Z : Type, In O' Z -> ExtendableAlong 2 (functor_coeq h k p q) (fun _ : Coeq f' g' => Z)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h
Z: Type
Z_inO: In O' Z

ExtendableAlong 2 (functor_coeq h k p q) (fun _ : Coeq f' g' => Z)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h
Z: Type
Z_inO: In O' Z

ExtendableAlong 2 k (fun _ : A' => Z)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h
Z: Type
Z_inO: In O' Z
forall u v : B' -> Z, ExtendableAlong 2 h (fun x : B' => u x = v x)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h
Z: Type
Z_inO: In O' Z

ExtendableAlong 2 k (fun _ : A' => Z)
nrapply (ooextendable_O_inverts O'); assumption.
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h
Z: Type
Z_inO: In O' Z

forall u v : B' -> Z, ExtendableAlong 2 h (fun x : B' => u x = v x)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A: Type
f, g: B -> A
B', A': Type
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
O_inverts0: O_inverts O' k
H0: IsConnMap O h
Z: Type
Z_inO: In O' Z
i:= inO_leq O' (Sep O): forall A : Type, In O' A -> In (Sep O) A

forall u v : B' -> Z, ExtendableAlong 2 h (fun x : B' => u x = v x)
intros u v; rapply (extendable_conn_map_inO O). Defined. (** And a similar property for pushouts *)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
O_inverts0: O_inverts O' k
O_inverts1: O_inverts O' l

O_inverts O' (functor_pushout h k l p q)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
O_inverts0: O_inverts O' k
O_inverts1: O_inverts O' l

O_inverts O' (functor_pushout h k l p q)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
O_inverts0: O_inverts O' k
O_inverts1: O_inverts O' l

O' <= Sep O
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
O_inverts0: O_inverts O' k
O_inverts1: O_inverts O' l
O_inverts O' (functor_sum k l)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
O_inverts0: O_inverts O' k
O_inverts1: O_inverts O' l
IsConnMap O h
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
O_inverts0: O_inverts O' k
O_inverts1: O_inverts O' l

O_inverts O' (functor_sum k l)
rapply O_inverts_functor_sum. Defined. (** And similar properties for connected maps *)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H0: IsConnMap O' k
H1: IsConnMap O h

IsConnMap O' (functor_coeq h k p q)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H0: IsConnMap O' k
H1: IsConnMap O h

IsConnMap O' (functor_coeq h k p q)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H0: IsConnMap O' k
H1: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O' (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)

ExtensionAlong (functor_coeq h k p q) P d
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H0: IsConnMap O' k
H1: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O' (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)

ExtendableAlong 1 k (fun x : A' => P (coeq x))
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H0: IsConnMap O' k
H1: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O' (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)
forall u v : forall x : B', P (coeq (g' x)), ExtendableAlong 1 h (fun x : B' => u x = v x)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H0: IsConnMap O' k
H1: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O' (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)

ExtendableAlong 1 k (fun x : A' => P (coeq x))
rapply ooextendable_conn_map_inO.
O, O': ReflectiveSubuniverse
H: O' <= Sep O
B, A, B', A': Type
f, g: B -> A
f', g': B' -> A'
h: B -> B'
k: A -> A'
p: k o f == f' o h
q: k o g == g' o h
H0: IsConnMap O' k
H1: IsConnMap O h
P: Coeq f' g' -> Type
P_inO: forall b : Coeq f' g', In O' (P b)
d: forall a : Coeq f g, P (functor_coeq h k p q a)

forall u v : forall x : B', P (coeq (g' x)), ExtendableAlong 1 h (fun x : B' => u x = v x)
pose (inO_leq O' (Sep O)); intros; rapply (ooextendable_conn_map_inO O). Defined.
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
H1: IsConnMap O' k
H2: IsConnMap O' l

IsConnMap O' (functor_pushout h k l p q)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
H1: IsConnMap O' k
H2: IsConnMap O' l

IsConnMap O' (functor_pushout h k l p q)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
H1: IsConnMap O' k
H2: IsConnMap O' l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O' (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)

ExtensionAlong (functor_pushout h k l p q) P d
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
H1: IsConnMap O' k
H2: IsConnMap O' l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O' (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)

ExtendableAlong 1 (functor_sum k l) (fun x : B' + C' => P (coeq x))
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
H1: IsConnMap O' k
H2: IsConnMap O' l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O' (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)
forall u v : forall x : A', P (coeq (inr (g' x))), ExtendableAlong 1 h (fun x : A' => u x = v x)
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
H1: IsConnMap O' k
H2: IsConnMap O' l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O' (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)

ExtendableAlong 1 (functor_sum k l) (fun x : B' + C' => P (coeq x))
apply extendable_functor_sum; rapply ooextendable_conn_map_inO.
O, O': ReflectiveSubuniverse
H: O' <= Sep O
A, B, C, A', B', C': Type
f: A -> B
g: A -> C
f': A' -> B'
g': A' -> C'
h: A -> A'
k: B -> B'
l: C -> C'
p: k o f == f' o h
q: l o g == g' o h
H0: IsConnMap O h
H1: IsConnMap O' k
H2: IsConnMap O' l
P: Pushout f' g' -> Type
P_inO: forall b : Pushout f' g', In O' (P b)
d: forall a : Pushout f g, P (functor_pushout h k l p q a)

forall u v : forall x : A', P (coeq (inr (g' x))), ExtendableAlong 1 h (fun x : A' => u x = v x)
pose (inO_leq O' (Sep O)); intros; rapply ooextendable_conn_map_inO. Defined. #[export] Hint Immediate inO_isequiv_to_O : typeclass_instances. #[export] Hint Immediate inO_unsigma : typeclass_instances. #[export] Hint Immediate isconnected_conn_map_to_unit : typeclass_instances.