Built with Alectryon, running Coq+SerAPI v8.19.0+0.19.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.
[Loading ML file number_string_notation_plugin.cmxs (using legacy method) ... done]
Require Import Truncations.Require Import ObjectClassifier Homotopy.ExactSequence Pointed.LocalOpen Scope type_scope.LocalOpen Scope path_scope.(** * BAut(X) *)(** ** Basics *)(** [BAut X] is the type of types that are merely equal to [X]. It is connected, by [is0connected_component] and any two points are merely equal by [merely_path_component]. *)DefinitionBAut (X : Type@{u}) := { Z : Type@{u} & merely (Z = X) }.CoercionBAut_pr1 X : BAut X -> Type := pr1.Instanceispointed_baut {X : Type} : IsPointed (BAut X) := (X; tr 1).(** We also define a pointed version [pBAut X], since the coercion [BAut_pr1] doesn't work if [BAut X] is a [pType]. *)DefinitionpBAut (X : Type) : pType
:= [BAut X, _].Definitionpath_baut `{Univalence} {X} (Z Z' : BAut X)
: (Z <~> Z') <~> (Z = Z' :> BAut X)
:= equiv_path_sigma_hprop _ _ oE equiv_path_universe _ _.
H: Univalence X: Type Z, Z': BAut X f: Z <~> Z'
ap (BAut_pr1 X) (path_baut Z Z' f) = path_universe f
H: Univalence X: Type Z, Z': BAut X f: Z <~> Z'
ap (BAut_pr1 X) (path_baut Z Z' f) = path_universe f
H: Univalence X: Type Z, Z': BAut X f: Z <~> Z'
ap pr1
(path_sigma_hprop Z Z' (path_universe_uncurried f)) =
path_universe f
apply ap_pr1_path_sigma_hprop.Defined.
H: Univalence X: Type Z, Z': BAut X f: Z <~> Z' z: Z
transport (funW : BAut X => W) (path_baut Z Z' f) z =
f z
H: Univalence X: Type Z, Z': BAut X f: Z <~> Z' z: Z
transport (funW : BAut X => W) (path_baut Z Z' f) z =
f z
H: Univalence X: Type Z, Z': BAut X f: Z <~> Z' z: Z
transport idmap (ap (BAut_pr1 X) (path_baut Z Z' f)) z =
f z
H: Univalence X: Type Z, Z': BAut X f: Z <~> Z' z: Z
transport idmap (ap (BAut_pr1 X) (path_baut Z Z' f)) z =
transport idmap (path_universe f) z
apply ap10, ap, ap_pr1_path_baut.Defined.(** The following tactic, which applies when trying to prove an hprop, replaces all assumed elements of [BAut X] by [X] itself. With [Univalence], this would work for any 0-connected type, but using [merely_path_component] we can avoid univalence. *)Ltacbaut_reduce :=
progressrepeatmatch goal with
| [ Z : BAut ?X |- _ ]
=> letZispoint := fresh"Zispoint"inassert (Zispoint := merely_path_component (point (BAut X)) Z);
revert Zispoint;
refine (@Trunc_ind _ _ _ _ _);
intro Zispoint;
destruct Zispoint
end.(** ** Truncation *)(** If [X] is an [n.+1]-type, then [BAut X] is an [n.+2]-type. *)
H: Univalence n: trunc_index X: Type IsTrunc0: IsTrunc n.+1 X
IsTrunc n.+2 (BAut X)
H: Univalence n: trunc_index X: Type IsTrunc0: IsTrunc n.+1 X
IsTrunc n.+2 (BAut X)
H: Univalence n: trunc_index X: Type IsTrunc0: IsTrunc n.+1 X
forallxy : BAut X, IsTrunc n.+1 (x = y)
H: Univalence n: trunc_index X: Type IsTrunc0: IsTrunc n.+1 X Z, W: BAut X
IsTrunc n.+1 (Z = W)
H: Univalence n: trunc_index X: Type IsTrunc0: IsTrunc n.+1 X
IsTrunc n.+1 (pt = pt)
exact (@istrunc_equiv_istrunc _ _ (path_baut _ _) n.+1 _).Defined.(** If [X] is truncated, then so is every element of [BAut X]. *)Instancetrunc_el_baut {nX} `{Funext} `{IsTrunc n X} (Z : BAut X)
: IsTrunc n Z
:= ltac:(by baut_reduce).(** ** Operations on [BAut] *)(** Multiplying by a fixed type *)Definitionbaut_prod_r (XA : Type)
: BAut X -> BAut (X * A)
:= funZ:BAut X =>
(Z * A ; Trunc_functor (-1) (ap (funW => W * A)) (pr2 Z))
: BAut (X * A).
H: Univalence X, A: Type Z, W: BAut X e: Z <~> W
ap (baut_prod_r X A) (path_baut Z W e) =
path_baut (baut_prod_r X A Z) (baut_prod_r X A W)
(equiv_functor_prod_r e)
H: Univalence X, A: Type Z, W: BAut X e: Z <~> W
ap (baut_prod_r X A) (path_baut Z W e) =
path_baut (baut_prod_r X A Z) (baut_prod_r X A W)
(equiv_functor_prod_r e)
H: Univalence X, A: Type Z, W: BAut X e: Z <~> W
ap (baut_prod_r X A)
(path_sigma_hprop Z W (path_universe_uncurried e)) =
path_sigma_hprop (baut_prod_r X A Z)
(baut_prod_r X A W)
(path_universe_uncurried (equiv_functor_prod_r e))
H: Univalence X, A: Type Z, W: BAut X e: Z <~> W
ap pr1
(ap (baut_prod_r X A)
(path_sigma_hprop Z W (path_universe_uncurried e))) =
path_universe_uncurried (equiv_functor_prod_r e)
H: Univalence X, A: Type Z, W: BAut X e: Z <~> W
ap (funx : BAut X => (baut_prod_r X A x).1)
(path_sigma_hprop Z W (path_universe_uncurried e)) =
path_universe_uncurried (equiv_functor_prod_r e)
H: Univalence X, A: Type Z, W: BAut X e: Z <~> W
ap (funZ : Type => Z * A)
(ap pr1
(path_sigma_hprop Z W (path_universe_uncurried e))) =
path_universe_uncurried (equiv_functor_prod_r e)
H: Univalence X, A: Type Z, W: BAut X e: Z <~> W
ap (funZ : Type => Z * A) (path_universe_uncurried e) =
path_universe_uncurried (equiv_functor_prod_r e)
H: Univalence X, A: Type Z, W: BAut X e: Z <~> W
equiv_path (Z * A) (W * A)
(ap (funZ : Type => Z * A)
(path_universe_uncurried e)) =
equiv_functor_prod_r e
apply ap_prod_r_path_universe.Qed.(** ** Centers *)(** The following lemma says that to define a section of a family [P] of hsets over [BAut X], it is equivalent to define an element of [P X] which is fixed by all automorphisms of [X]. *)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
{e : P pt &
forallg : X <~> X,
transport P (path_universe g) e = e} <~>
(forallZ : BAut X, P Z)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
{e : P pt &
forallg : X <~> X,
transport P (path_universe g) e = e} <~>
(forallZ : BAut X, P Z)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
{e : P pt &
forallg : X <~> X,
transport P (path_universe g) e = e} <~>
(forall (x : Type) (y : merely (x = X)), P (x; y))
(** We use the fact that maps out of a propositional truncation into an hset are equivalent to weakly constant functions. *)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) Z: Type
Z = X -> IsHSet (P Z)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
{e : P X &
forallg : X <~> X,
transport P (path_universe g) e = e} <~>
(foralla : Type,
{f : a = X -> P a & WeaklyConstant f})
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) Z: Type
Z = X -> IsHSet (P Z)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) Z: Type p: Z = X
IsHSet (P Z)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) Z: Type p: Z = X
IsHSet (P (Z; tr p))
exact _.
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
{e : P X &
forallg : X <~> X,
transport P (path_universe g) e = e} <~>
(foralla : Type,
{f : a = X -> P a & WeaklyConstant f})
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
{e : P X &
forallg : X <~> X,
transport P (path_universe g) e = e} <~>
(foralla : Type,
{f : a = X -> P a & forallxy : a = X, f x = f y})
(** Now we peel away a bunch of contractible types. *)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
{e : P X &
forallg : X <~> X,
transport P (path_universe g) e = e} <~>
{f : forallx : Type, x = X -> P x &
forall (x : Type) (x0y : x = X), f x x0 = f x y}
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
P X <~> (forallx : Type, x = X -> P x)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
foralla : P X,
(fune : P X =>
forallg : X <~> X,
transport P (path_universe g) e = e) a <~>
(funf : forallx : Type, x = X -> P x =>
forall (x : Type) (x0y : x = X), f x x0 = f x y)
(?f a)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z)
foralla : P X,
(fune : P X =>
forallg : X <~> X,
transport P (path_universe g) e = e) a <~>
(funf : forallx : Type, x = X -> P x =>
forall (x : Type) (x0y : x = X), f x x0 = f x y)
(equiv_paths_ind_r X
(fun (x : Type) (_ : x = X) => P x) a)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X
(forallg : X <~> X,
transport P (path_universe g) p = p) <~>
(forall (x : Type) (x0y : x = X),
paths_ind_r X (fun (x1 : Type) (_ : x1 = X) => P x1)
p x x0 =
paths_ind_r X (fun (x1 : Type) (_ : x1 = X) => P x1)
p x y)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X
(forallg : X <~> X,
transport P (path_universe g) p = p) <~>
(forally : X = X,
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
1 =
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
y)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X
X = X <~> (X <~> X)
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X
forallb : X = X,
(funa : X <~> X =>
transport P (path_universe a) p = p) (?f b) <~>
(funb0 : X = X =>
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
1 =
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
b0) b
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X
forallb : X = X,
(funa : X <~> X =>
transport P (path_universe a) p = p)
(equiv_equiv_path X X b) <~>
(funb0 : X = X =>
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
1 =
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
b0) b
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X e: X = X
transport P (path_universe (transport idmap e)) p = p <~>
p =
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
e
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X e: X = X
p = transport P (path_universe (transport idmap e))^ p <~>
p =
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
e
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X e: X = X
transport P (path_universe (transport idmap e))^ p =
paths_ind_r X (fun (x : Type) (_ : x = X) => P x) p X
e
H: Univalence X: Type P: Type -> Type H0: forallZ : BAut X, IsHSet (P Z) p: P X e: X = X
transport P e^ p = transport P e^ p
reflexivity.Defined.(** This implies that if [X] is a set, then the center of [BAut X] is the set of automorphisms of [X] that commute with every other automorphism (i.e. the center, in the usual sense, of the group of automorphisms of [X]). *)
H: Univalence X: Type IsHSet0: IsHSet X
{f : X <~> X & forallg : X <~> X, g o f == f o g} <~>
(forallZ : BAut X, Z = Z)
H: Univalence X: Type IsHSet0: IsHSet X
{f : X <~> X & forallg : X <~> X, g o f == f o g} <~>
(forallZ : BAut X, Z = Z)
H: Univalence X: Type IsHSet0: IsHSet X
{f : X <~> X &
forallg : X <~> X,
(funx : X => g (f x)) == (funx : X => f (g x))} <~>
(foralla : {Z : Type & merely (Z = X)}, a.1 = a.1)
H: Univalence X: Type IsHSet0: IsHSet X
{f : X <~> X &
forallg : X <~> X,
(funx : X => g (f x)) == (funx : X => f (g x))} <~>
{e : pt = pt &
forallg : X <~> X,
transport (funZ : Type => Z = Z) (path_universe g) e =
e}
H: Univalence X: Type IsHSet0: IsHSet X
{f : X <~> X &
forallg : X <~> X,
(funx : X => g (f x)) == (funx : X => f (g x))} <~>
{e : X = X &
forallg : X <~> X,
transport (funZ : Type => Z = Z) (path_universe g) e =
e}
H: Univalence X: Type IsHSet0: IsHSet X f: X <~> X
(forallg : X <~> X,
(funx : X => g (f x)) == (funx : X => f (g x))) <~>
(forallg : X <~> X,
transport (funZ : Type => Z = Z) (path_universe g)
(equiv_path_universe X X f) =
equiv_path_universe X X f)
H: Univalence X: Type IsHSet0: IsHSet X f, g: X <~> X
(funx : X => g (f x)) == (funx : X => f (g x)) <~>
transport (funZ : Type => Z = Z) (path_universe g)
(path_universe_uncurried f) =
path_universe_uncurried f
H: Univalence X: Type IsHSet0: IsHSet X f, g: X <~> X
(funx : X => g (f x)) = (funx : X => f (g x)) <~>
transport (funZ : Type => Z = Z) (path_universe g)
(path_universe_uncurried f) =
path_universe_uncurried f
H: Univalence X: Type IsHSet0: IsHSet X f, g: X <~> X
g oE f = f oE g <~>
transport (funZ : Type => Z = Z) (path_universe g)
(path_universe_uncurried f) =
path_universe_uncurried f
H: Univalence X: Type IsHSet0: IsHSet X f: X <~> X
forallg : X <~> X,
g oE f = f oE g <~>
transport (funZ : Type => Z = Z) (path_universe g)
(path_universe_uncurried f) =
path_universe_uncurried f
H: Univalence X: Type IsHSet0: IsHSet X f: X <~> X g: X = X
equiv_path X X g oE f = f oE equiv_path X X g <~>
transport (funZ : Type => Z = Z)
(path_universe (equiv_path X X g))
(path_universe_uncurried f) =
path_universe_uncurried f
H: Univalence X: Type IsHSet0: IsHSet X g: X = X
forallf : X <~> X,
equiv_path X X g oE f = f oE equiv_path X X g <~>
transport (funZ : Type => Z = Z)
(path_universe (equiv_path X X g))
(path_universe_uncurried f) =
path_universe_uncurried f
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
equiv_path X X g oE equiv_path X X f =
equiv_path X X f oE equiv_path X X g <~>
transport (funZ : Type => Z = Z)
(path_universe (equiv_path X X g))
(path_universe_uncurried (equiv_path X X f)) =
path_universe_uncurried (equiv_path X X f)
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
equiv_path X X (f @ g) =
equiv_path X X f oE equiv_path X X g <~>
transport (funZ : Type => Z = Z)
(path_universe (equiv_path X X g))
(path_universe_uncurried (equiv_path X X f)) =
path_universe_uncurried (equiv_path X X f)
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
equiv_path X X (f @ g) = equiv_path X X (g @ f) <~>
transport (funZ : Type => Z = Z)
(path_universe (equiv_path X X g))
(path_universe_uncurried (equiv_path X X f)) =
path_universe_uncurried (equiv_path X X f)
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
f @ g = g @ f <~>
transport (funZ : Type => Z = Z)
(path_universe (equiv_path X X g))
(path_universe_uncurried (equiv_path X X f)) =
path_universe_uncurried (equiv_path X X f)
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
f @ g = g @ f <~>
((path_universe (equiv_path X X g))^ @
path_universe_uncurried (equiv_path X X f)) @
path_universe (equiv_path X X g) =
path_universe_uncurried (equiv_path X X f)
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
f @ g = g @ f <~>
(path_universe (equiv_path X X g))^ @
(path_universe_uncurried (equiv_path X X f) @
path_universe (equiv_path X X g)) =
path_universe_uncurried (equiv_path X X f)
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
f @ g = g @ f <~>
path_universe_uncurried (equiv_path X X f) @
path_universe (equiv_path X X g) =
path_universe (equiv_path X X g) @
path_universe_uncurried (equiv_path X X f)
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
path_universe_uncurried (equiv_path X X f) @
path_universe (equiv_path X X g) = f @ g
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
g @ f =
path_universe (equiv_path X X g) @
path_universe_uncurried (equiv_path X X f)
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
path_universe_uncurried (equiv_path X X f) @
path_universe (equiv_path X X g) = f @ g
apply concat2; apply eissect.
H: Univalence X: Type IsHSet0: IsHSet X g, f: X = X
g @ f =
path_universe (equiv_path X X g) @
path_universe_uncurried (equiv_path X X f)
symmetry; apply concat2; apply eissect.Defined.(** We show that this equivalence takes the identity equivalence to the identity in the center. We have to be careful in this proof never to [simpl] or [unfold] too many things, or Coq will produce gigantic terms that take it forever to compute with. *)
H: Univalence X: Type IsHSet0: IsHSet X
center_baut X
(1%equiv; fun (g : X <~> X) (x : X) => 1) =
(funZ : BAut X => 1)
H: Univalence X: Type IsHSet0: IsHSet X
center_baut X
(1%equiv; fun (g : X <~> X) (x : X) => 1) =
(funZ : BAut X => 1)
H: Univalence X: Type IsHSet0: IsHSet X Z: BAut X
center_baut X
(1%equiv; fun (g : X <~> X) (x : X) => 1) Z = 1
H: Univalence X: Type IsHSet0: IsHSet X Z: BAut X X0: IsHSet (Z.1 = Z.1)
center_baut X
(1%equiv; fun (g : X <~> X) (x : X) => 1) Z = 1
H: Univalence X: Type IsHSet0: IsHSet X X0: IsHSet (pt.1 = pt.1)
center_baut X
(1%equiv; fun (g : X <~> X) (x : X) => 1) pt = 1
exact (ap (path_sigma_hprop _ _) path_universe_1
@ path_sigma_hprop_1 _).Defined.(** Similarly, if [X] is a 1-type, we can characterize the 2-center of [BAut X]. *)(** Coq is too eager about unfolding some things appearing in this proof. *)SectionCenter2BAut.LocalArguments equiv_path_equiv : simpl never.LocalArguments equiv_path2_universe : simpl never.
H: Univalence X: Type IsTrunc0: IsTrunc 1 X
{f : forallx : X, x = x &
forall (g : X <~> X) (x : X), ap g (f x) = f (g x)} <~>
(forallZ : BAut X, 1 = 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X
{f : forallx : X, x = x &
forall (g : X <~> X) (x : X), ap g (f x) = f (g x)} <~>
(forallZ : BAut X, 1 = 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X Z: {Z : Type & merely (Z = X)}
1 = equiv_path_sigma_hprop Z Z 1
H: Univalence X: Type IsTrunc0: IsTrunc 1 X Z: {Z : Type & merely (Z = X)}
equiv_path_sigma_hprop Z Z 1 = 1
H: Univalence X: Type IsTrunc0: IsTrunc 1 X
{f : forallx : X, x = x &
forall (g : X <~> X) (x : X), ap g (f x) = f (g x)} <~>
(foralla : {Z : Type & merely (Z = X)}, 1 = 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X Z: {Z : Type & merely (Z = X)}
1 = equiv_path_sigma_hprop Z Z 1
symmetry; apply path_sigma_hprop_1.
H: Univalence X: Type IsTrunc0: IsTrunc 1 X Z: {Z : Type & merely (Z = X)}
equiv_path_sigma_hprop Z Z 1 = 1
H: Univalence X: Type IsTrunc0: IsTrunc 1 X
{f : forallx : X, x = x &
forall (g : X <~> X) (x : X), ap g (f x) = f (g x)} <~>
(foralla : {Z : Type & merely (Z = X)}, 1 = 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X Z: {Z : Type & merely (Z = X)}
equiv_path_sigma_hprop Z Z 1 = 1
apply path_sigma_hprop_1.
H: Univalence X: Type IsTrunc0: IsTrunc 1 X
{f : forallx : X, x = x &
forall (g : X <~> X) (x : X), ap g (f x) = f (g x)} <~>
(foralla : {Z : Type & merely (Z = X)}, 1 = 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: forallZ : BAut X, IsHSet (1 = 1)
{f : forallx : X, x = x &
forall (g : X <~> X) (x : X), ap g (f x) = f (g x)} <~>
(foralla : {Z : Type & merely (Z = X)}, 1 = 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: forallZ : BAut X, IsHSet (1 = 1)
{f : forallx : X, x = x &
forall (g : X <~> X) (x : X), ap g (f x) = f (g x)} <~>
{e : 1 = 1 &
forallg : X <~> X,
transport (funZ : Type => 1 = 1) (path_universe g) e =
e}
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: forallZ : BAut X, IsHSet (1 = 1)
(forallx : X, x = x) <~> 1 = 1
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: forallZ : BAut X, IsHSet (1 = 1)
foralla : forallx : X, x = x,
(funf : forallx : X, x = x =>
forall (g : X <~> X) (x : X), ap g (f x) = f (g x)) a <~>
(fune : 1 = 1 =>
forallg : X <~> X,
transport (funZ : Type => 1 = 1) (path_universe g) e =
e) (?f a)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: forallZ : BAut X, IsHSet (1 = 1)
(forallx : X, x = x) <~> 1 = 1
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: forallZ : BAut X, IsHSet (1 = 1)
abstract (rewrite !whiskerL_pp, !concat_pp_p; reflexivity).Defined.(** Once again we compute it on the identity. In this case it seems to be unavoidable to do some [simpl]ing (or at least [cbn]ing), making this proof somewhat slower. *)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X
center2_baut X
(funx : X => 1; fun (g : X <~> X) (x : X) => 1) =
(funZ : BAut X => 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X
center2_baut X
(funx : X => 1; fun (g : X <~> X) (x : X) => 1) =
(funZ : BAut X => 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X Z: BAut X
center2_baut X
(funx : X => 1; fun (g : X <~> X) (x : X) => 1) Z =
1
H: Univalence X: Type IsTrunc0: IsTrunc 1 X Z: BAut X X0: IsHSet (1 = 1)
center2_baut X
(funx : X => 1; fun (g : X <~> X) (x : X) => 1) Z =
1
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: IsHSet (1 = 1)
center2_baut X
(funx : X => 1; fun (g : X <~> X) (x : X) => 1) pt =
1
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: IsHSet (1 = 1)
H: Univalence X: Type IsTrunc0: IsTrunc 1 X X0: IsHSet (1 = 1)
1 = 1
reflexivity.Defined.EndCenter2BAut.SectionClassifyingMaps.(** ** Maps into [BAut F] classify bundles with fiber [F] *)(** The property of being merely equivalent to a given type [F] defines a subuniverse. *)
F: Type
Subuniverse
F: Type
Subuniverse
F: Type
forallTU : Type,
merely (T <~> F) ->
forallf : T -> U, IsEquiv f -> merely (U <~> F)
F, T, U: Type mere_eq: merely (T <~> F) f: T -> U iseq_f: IsEquiv f
merely (U <~> F)
F, T, U: Type f: T -> U iseq_f: IsEquiv f mere_eq: T <~> F
merely (U <~> F)
F, T, U: Type f: T -> U iseq_f: IsEquiv f mere_eq: T <~> F feq:= {| equiv_fun := f; equiv_isequiv := iseq_f |}: T <~> U
merely (U <~> F)
exact (tr (mere_eq oE feq^-1)).Defined.(** The universe of O-local types for [subuniverse_merely_equiv F] is equivalent to [BAut F]. *)
H: Univalence F: Type
BAut F <~> Type_ (subuniverse_merely_equiv F)
H: Univalence F: Type
BAut F <~> Type_ (subuniverse_merely_equiv F)
H: Univalence F, X: Type
Trunc (-1) (X = F) <~> Trunc (-1) (X <~> F)
H: Univalence F, X: Type
X = F <~> (X <~> F)
exact (equiv_path_universe _ _)^-1%equiv.Defined.(** Consequently, maps into [BAut F] correspond to bundles with fibers merely equivalent to [F]. *)
H: Univalence Y: pType F: Type
(Y -> BAut F) <~>
{p : Slice Y &
forally : Y, merely (hfiber p.2 y <~> F)}
H: Univalence Y: pType F: Type
(Y -> BAut F) <~>
{p : Slice Y &
forally : Y, merely (hfiber p.2 y <~> F)}
H: Univalence Y: pType F: Type
(Y -> Type_ (subuniverse_merely_equiv F)) <~>
{p : Slice Y &
forally : Y, merely (hfiber p.2 y <~> F)}
H: Univalence Y: pType F: Type
{p : {X : Type & X -> Y} &
MapIn (subuniverse_merely_equiv F) p.2} <~>
{p : Slice Y &
forally : Y, merely (hfiber p.2 y <~> F)}
H: Univalence Y: pType F: Type p: {X : Type & X -> Y}
(funp : {X : Type & X -> Y} =>
MapIn (subuniverse_merely_equiv F) p.2) p <~>
(funp : Slice Y =>
forally : Y, merely (hfiber p.2 y <~> F)) p
H: Univalence Y: pType F: Type p: {X : Type & X -> Y} y: Y
In (subuniverse_merely_equiv F) (hfiber p.2 y) <~>
merely (hfiber p.2 y <~> F)
byapply Trunc_functor_equiv.Defined.(** The pointed version of [equiv_baut_typeO] above. *)
H: Univalence F: Type
pBAut F <~>*
[Type_ (subuniverse_merely_equiv F), (F; tr 1%equiv)]
H: Univalence F: Type
pBAut F <~>*
[Type_ (subuniverse_merely_equiv F), (F; tr 1%equiv)]
H: Univalence F: Type
BAut F <~> Type_ (subuniverse_merely_equiv F)
H: Univalence F: Type
?Goal ispointed_baut = (F; tr 1%equiv)
H: Univalence F: Type
equiv_baut_typeO ispointed_baut = (F; tr 1%equiv)
byapply path_sigma_hprop.Defined.Definitionequiv_pmap_pbaut_pfibration `{Univalence} {Y F : pType@{u}}
: (Y ->* pBAut@{u v} F) <~> { p : { q : pSlice Y & forally:Y, merely (hfiber q.2 y <~> F) } &
pfiber p.1.2 <~>* F }
:= (equiv_sigma_pfibration_O (subuniverse_merely_equiv F))
oE pequiv_pequiv_postcompose pequiv_pbaut_typeOp.(** When [Y] is connected, pointed maps into [pBAut F] correspond to maps into the universe sending the base point to [F]. *)
H: Univalence Y: pType F: Type H0: IsConnected (Tr 0) Y
(Y ->* pBAut F) <~> (Y ->* [Type, F])
H: Univalence Y: pType F: Type H0: IsConnected (Tr 0) Y
(Y ->* pBAut F) <~> (Y ->* [Type, F])
H: Univalence Y: pType F: Type H0: IsConnected (Tr 0) Y
exact equiv_pmap_typeO_type_connected.Defined.(** When [Y] is connected, [pBAut F] classifies fiber sequences over [Y] with fiber [F]. *)Definitionequiv_pmap_pbaut_pfibration_connected `{Univalence} {Y F : pType} `{IsConnected 0 Y}
: (Y ->* pBAut F) <~> { X : pType & FiberSeq F X Y }
:= classify_fiberseq oE equiv_pmap_pbaut_type_p.EndClassifyingMaps.