Built with Alectryon, running Coq+SerAPI v8.20.0+0.20.0. 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.
(** * Univalent Foundations, Part AVladimir Voevodsky.Feb. 2010 - Sep. 2011.This file is based on the first part of the original uu0 file.The uu0 file contained the basic results of the univalent foundationsthat required the use of only one universe.Eventually the requirement concerning one universe was removed because of thegeneral uncertainty in what does it mean for a construction to require only one universe.For example, [ boolsumfun ], when written in terms of the eliminatior [ bool_rect ]instead of the [ match ], requires application of [ bool_rect ] to an argument that isnot a member of the base universe [ UU ]. This would be different if the universe managementin Coq was constructed differently. Due to this uncertainty we do not consider any more thesingle universe requirement as a defining one when selecting results for the inclusion in Foundations.Part A was created as a separate file on Dec. 3, 2014.Together with Part B it contains those results that do not require any axioms.This file was edited and expanded by Benedikt Ahrens 2014-2016, Dan Grayson 2014-2016,Vladimir Voevodsky 2014-2016, Alex Kavvos 2014, Peter LeFanu Lumsdaine 2016 andTomi Pannila 2016.*)(** ** Contents- Preamble - Settings - Imports- Some standard constructions not using identity types (paths) - Canonical functions from [ empty ] and to [unit ] - Identity functions and function composition, curry and uncurry - Iteration of an endomorphism - Basic constructions related to the adjoint evaluation function [ X -> ((X -> Y) -> Y) ] - Pairwise direct products - Negation and double negation - Logical equivalence- Paths and operations on paths - Associativity of function composition and mutual invertibility of curry/uncurry - Composition of paths and inverse paths - Direct product of paths - The function [ maponpaths ] between paths types defined by a function between ambient types - [ maponpaths ] for the identity functions and compositions of functions - Homotopy between functions - Equality between functions defines a homotopy - [ maponpaths ] for a function homotopic to the identity - [ maponpaths ] in the case of a projection p with a section s - Fibrations and paths - the transport functions - A series of lemmas about paths and [ total2 ] - Lemmas about transport adapted from the HoTT library and the HoTT book - Homotopies between families and the total spaces- First fundamental notions - Contractibility [ iscontr ] - Homotopy fibers [ hfiber ] - The functions between the hfibers of homotopic functions over the same point - Paths in homotopy fibers - Coconuses: spaces of paths that begin (coconusfromt) or end (coconustot) at a given point - The total paths space of a type - two definitions - Coconus of a function: the total space of the family of h-fibers- Weak equivalences - Basics - [ isweq ] and [ weq ] - Weak equivalences and paths spaces (more results in further sections) - Adjointness property of a weak equivalence and its inverse - Transport functions are weak equivalences - Weak equivalences between contractible types (one implication) - Unit and contractibility - Homotopy equivalence is a weak equivalence - Some weak equivalences - 2-out-of-3 property of weak equivalences - Any function between contractible types is a weak equivalence - Composition of weak equivalences - 2-out-of-6 property of weak equivalences - Pairwise direct products of weak equivalences - Weak equivalence of a type and its direct product with the unit - Associativity of total2 as a weak equivalence - Associativity and commutativity of direct products as weak equivalences- Binary coproducts and their basic properties - Distributivity of coproducts and direct products as a weak equivalence - Total space of a family over a coproduct - Pairwise sum of functions, coproduct associativity and commutativity - Coproduct with a "negative" type - Coproduct of two functions - The [ equality_cases ] construction and four applications to [ ii1 ] and [ ii2 ] - Bool as coproduct - Pairwise coproducts as dependent sums of families over [ bool ] - Splitting of [ X ] into a coproduct defined by a function [ X -> Y ⨿ Z ] - Some properties of bool - Fibrations with only one non-empty fiber- Basics about fibration sequences - The structures of a complex and of a fibration sequence on a composable pair of functions - Construction of the derived fibration sequence - Explicit description of the first map in the second derived sequence - Fibration sequences based on [ tpair P z ] and [ pr1 : total2 P -> Z ] ( the "pr1-case" ) - Fibration sequences based on [ hfiberpr1 : hfiber g z -> Y ] and [ g : Y -> Z ] (the "g-case") - Fibration sequence of h-fibers defined by a composable pair of functions (the "hf-case")- Functions between total spaces of families - Function [ totalfun ] between total spaces from a family of functions between the fibers - Function [ fpmap ] between the total spaces from a function between the bases - Homotopy fibers of [ fpmap ] - The [ fpmap ] from a weak equivalence is a weak equivalence - Total spaces of families over a contractible base - Function on the total spaces from functions between the bases and between the fibers- Homotopy fiber squares - Homotopy commutative squares - Short complexes and homotopy commutative squares - Homotopy fiber products - Homotopy fiber products and homotopy fibers - Homotopy fiber squares - Fiber sequences and homotopy fiber squares *)(** ** Preamble *)(** *** Imports *)Require Export UniMath.Foundations.Preamble.(* end of "Preamble" *)(** ** Some standard constructions not using identity types (paths) *)(** *** Canonical functions from [ empty ] and to [ unit ] *)
∏ X : UU, ∅ → X
(* type this in emacs in agda-input methodwith \prod *)
∏ X : UU, ∅ → X
X: UU
∅ → X
X: UU H: ∅
X
induction H.Defined.Arguments fromempty { X } _.Definitiontounit {X : UU} : X -> unit := λ_, tt.(** *** Functions from [ unit ] corresponding to terms *)Definitiontermfun {X : UU} (x : X) : unit -> X := λ_, x.(** *** Identity functions and function composition, curry and uncurry *)Definitionidfun (T : UU) := λt:T, t.(** makes [simpl], [cbn], etc. unfold [idfun X x] but not [ idfun X ]: *)Arguments idfun _ _ /.Definitionfuncomp {XY : UU} {Z:Y->UU} (f : X -> Y) (g : ∏ y:Y, Z y) := λx, g (f x).(** make [simpl], [cbn], etc. unfold [ (f ∘ g) x ] but not [ f ∘ g ]: *)Arguments funcomp {_ _ _} _ _ _/.Declare Scope functions.Delimit Scope functions with functions.Open Scope functions.Notation"g ∘ f" := (funcomp f g) : functions.(** back and forth between functions of pairs and functions returning functions *)Definitioncurry {X : UU} {Y : X -> UU} {Z : (∑ x, Y x) -> UU}
(f : ∏ p, Z p) : (∏ x : X, ∏ y : Y x, Z (x,, y)) :=
λxy, f (x,, y).Definitionuncurry {X : UU} {Y : X -> UU} {Z : (∑ x, Y x) -> UU}
(g : ∏ x : X, ∏ y : Y x, Z (x,, y)) : (∏ p, Z p) :=
λx, g (pr1 x) (pr2 x).(** *** Definition of binary operation *)Definitionbinop (X : UU) : UU := X -> X -> X.(** *** Iteration of an endomorphism *)
T: UU f: T → T n: nat
T → T
T: UU f: T → T n: nat
T → T
T: UU f: T → T
T → T
T: UU f: T → T n: nat IHn: T → T
T → T
T: UU f: T → T
T → T
exact (idfun T).
T: UU f: T → T n: nat IHn: T → T
T → T
exact (f ∘ IHn).Defined.(** *** Basic constructions related to the adjoint evaluation function [ X -> ((X -> Y) -> Y) ] *)Definitionadjev {XY : UU} (x : X) (f : X -> Y) : Y := f x.Definitionadjev2 {XY : UU} (phi : ((X -> Y) -> Y) -> Y) : X -> Y :=
λx, phi (λf, f x).(** *** Pairwise direct products *)Definitiondirprod (XY : UU) := ∑ x:X, Y.Notation"A × B" := (dirprod A B) : type_scope.Definitiondirprod_pr1 {XY : UU} := pr1 : X × Y -> X.Definitiondirprod_pr2 {XY : UU} := pr2 : X × Y -> Y.Definitionmake_dirprod {XY : UU} (x:X) (y:Y) : X × Y := x,,y.Definitiondirprodadj {XYZ : UU} (f : X × Y -> Z) : X -> Y -> Z :=
λxy, f (x,,y).Definitiondirprodf {XYX'Y' : UU}
(f : X -> Y) (f' : X' -> Y') (xx' : X × X') : Y × Y' :=
make_dirprod (f (pr1 xx')) (f' (pr2 xx')).
X, Y, P: UU xp: (X → P) → P yp: (Y → P) → P
(X × Y → P) → P
X, Y, P: UU xp: (X → P) → P yp: (Y → P) → P
(X × Y → P) → P
X, Y, P: UU xp: (X → P) → P yp: (Y → P) → P X0: X × Y → P
P
X, Y, P: UU xp: (X → P) → P yp: (Y → P) → P X0: X × Y → P
X → P
X, Y, P: UU xp: (X → P) → P yp: (Y → P) → P X0: X × Y → P x: X
P
X, Y, P: UU xp: (X → P) → P yp: (Y → P) → P X0: X × Y → P x: X
Y → P
X, Y, P: UU xp: (X → P) → P yp: (Y → P) → P X0: X × Y → P x: X y: Y
P
apply (X0 (make_dirprod x y)).Defined.(** *** Negation and double negation *)Definitionneg (X : UU) : UU := X -> empty.Notation"'¬' X" := (neg X).(* type this in emacs in agda-input method with \neg *)Notation"x != y" := (neg (x = y)) : type_scope.Definitionnegf {XY : UU} (f : X -> Y) : ¬ Y -> ¬ X := λphix, phi (f x).Definitiondneg (X : UU) : UU := ¬ ¬ X.Notation"'¬¬' X" := (dneg X).(* type this in emacs in agda-input method with \neg twice *)Definitiondnegf {XY : UU} (f : X -> Y) : dneg X -> dneg Y :=
negf (negf f).Definitiontodneg (X : UU) : X -> dneg X := adjev.Definitiondnegnegtoneg {X : UU} : ¬¬ ¬ X -> ¬ X := adjev2.
exact (pr1 j ∘ pr1 i,, pr2 i ∘ pr2 j).Defined.(* end of "Some standard constructions not using identity types (paths)". *)(** ** Paths and operations on [ paths ] *)(** *** Associativity of function composition and mutual invertibility of curry/uncurry *)(** While the paths in two of the three following lemmas are trivial, having them aslemmas turns out to be convenient in some future proofs. They are used to apply a particulardefinitional equality to modify the syntactic form of the goal in order to make thenext tactic, which uses the syntactic form of the goal to guess how to proceed, to work.The same applies to other lemmas below whose proof is by immediate "reflexivity" or"idpath". *)
X, Y, Z, W: UU f: X → Y g: Y → Z h: Z → W
h ∘ (g ∘ f) = h ∘ g ∘ f
X, Y, Z, W: UU f: X → Y g: Y → Z h: Z → W
h ∘ (g ∘ f) = h ∘ g ∘ f
X, Y, Z, W: UU f: X → Y g: Y → Z h: Z → W
h ∘ (g ∘ f) = h ∘ g ∘ f
apply idpath.Defined.
X, Z: UU Y: X → UU f: (∑ x : X, Y x) → Z
∏ p : ∑ x : X, Y x, uncurry (curry f) p = f p
X, Z: UU Y: X → UU f: (∑ x : X, Y x) → Z
∏ p : ∑ x : X, Y x, uncurry (curry f) p = f p
X, Z: UU Y: X → UU f: (∑ x : X, Y x) → Z p: ∑ x : X, Y x
uncurry (curry f) p = f p
X, Z: UU Y: X → UU f: (∑ x : X, Y x) → Z x: X y: Y x
uncurry (curry f) (x,, y) = f (x,, y)
apply idpath.Defined.
X, Z: UU Y: X → UU g: ∏ x : X, Y x → Z
∏ (x : X) (y : Y x), curry (uncurry g) x y = g x y
X, Z: UU Y: X → UU g: ∏ x : X, Y x → Z
∏ (x : X) (y : Y x), curry (uncurry g) x y = g x y
X, Z: UU Y: X → UU g: ∏ x : X, Y x → Z x: X y: Y x
curry (uncurry g) x y = g x y
apply idpath.Defined.(** *** Composition of paths and inverse paths *)
X: UU a, b, c: X e1: a = b e2: b = c
a = c
X: UU a, b, c: X e1: a = b e2: b = c
a = c
X: UU a, b, c: X e1: a = b e2: b = c
a = c
X: UU a, c: X e2: a = c
a = c
apply e2.Defined.#[global]
Hint Resolve @pathscomp0 : pathshints.Ltacintermediate_path x := apply (pathscomp0 (b := x)).Ltacetrans := eapply pathscomp0.Notation"p @ q" := (pathscomp0 p q).
X: UU a, b: X e1: a = b
e1 @ idpath b = e1
X: UU a, b: X e1: a = b
e1 @ idpath b = e1
X: UU a, b: X e1: a = b
e1 @ idpath b = e1
X: UU a: X
idpath a @ idpath a = idpath a
X: UU a: X
idpath a = idpath a
apply idpath.Defined.(** Note that we do not introduce [ pathscomp0lid ] since the corresponding two terms are convertible to each other due to our definition of [ pathscomp0 ]. If we defined it by inductioning [ e2 ] and applying [ e1 ] then [ pathscomp0rid ] would be trivial but [ pathscomp0lid ] would require a proof. Similarly we do not introduce a lemma to connect [ pathsinv0 (idpath _) ] to [ idpath ]. *)
X: UU a, b: X e: a = b
b = a
X: UU a, b: X e: a = b
b = a
X: UU a, b: X e: a = b
b = a
X: UU a: X
a = a
apply idpath.Defined.#[global]
Hint Resolve @pathsinv0 : pathshints.
X: Type a, b, c, d: X f: a = b g: b = c h: c = d
f @ g @ h = (f @ g) @ h
X: Type a, b, c, d: X f: a = b g: b = c h: c = d
f @ g @ h = (f @ g) @ h
X: Type a, b, c, d: X f: a = b g: b = c h: c = d
f @ g @ h = (f @ g) @ h
X: Type a, c, d: X g: a = c h: c = d
idpath a @ g @ h = (idpath a @ g) @ h
apply idpath.Defined.Notation"! p " := (pathsinv0 p).
X: UU a, b: X e: a = b
! e @ e = idpath b
X: UU a, b: X e: a = b
! e @ e = idpath b
X: UU a, b: X e: a = b
! e @ e = idpath b
X: UU a: X
! idpath a @ idpath a = idpath a
apply idpath.Defined.
X: UU a, b: X e: a = b
e @ ! e = idpath a
X: UU a, b: X e: a = b
e @ ! e = idpath a
X: UU a, b: X e: a = b
e @ ! e = idpath a
X: UU a: X
idpath a @ ! idpath a = idpath a
apply idpath.Defined.
X: UU x, x': X e: x = x'
! (! e) = e
X: UU x, x': X e: x = x'
! (! e) = e
X: UU x, x': X e: x = x'
! (! e) = e
X: UU x: X
! (! idpath x) = idpath x
apply idpath.Defined.
X: UU x, y, z: X p: x = y r, s: y = z
p @ r = p @ s → r = s
X: UU x, y, z: X p: x = y r, s: y = z
p @ r = p @ s → r = s
X: UU x, y, z: X p: x = y r, s: y = z e: p @ r = p @ s
r = s
X: UU x, z: X r, s: x = z e: idpath x @ r = idpath x @ s
r = s
exact e.Defined.
X: UU x, y, z: X p, q: x = y s: y = z
p @ s = q @ s → p = q
X: UU x, y, z: X p, q: x = y s: y = z
p @ s = q @ s → p = q
X: UU x, y, z: X p, q: x = y s: y = z e: p @ s = q @ s
p = q
X: UU x, y: X p, q: x = y e: p @ idpath y = q @ idpath y
p = q
X: UU x, y: X p, q: x = y e: p @ idpath y = q @ idpath y
p = p @ idpath y
X: UU x, y: X p, q: x = y e: p @ idpath y = q @ idpath y
q @ idpath y = q
X: UU x, y: X p, q: x = y e: p @ idpath y = q @ idpath y
p = p @ idpath y
apply pathsinv0, pathscomp0rid.
X: UU x, y: X p, q: x = y e: p @ idpath y = q @ idpath y
q @ idpath y = q
apply pathscomp0rid.Defined.
X: UU x, y, z: X p: x = y q: y = z
! (p @ q) = ! q @ ! p
X: UU x, y, z: X p: x = y q: y = z
! (p @ q) = ! q @ ! p
X: UU x, z: X q: x = z
! (idpath x @ q) = ! q @ ! idpath x
X: UU x: X
! (idpath x @ idpath x) = ! idpath x @ ! idpath x
apply idpath.Defined.(** *** Direct product of paths *)
X, Y: UU x1, x2: X y1, y2: Y ex: x1 = x2 ey: y1 = y2
make_dirprod x1 y1 = make_dirprod x2 y2
X, Y: UU x1, x2: X y1, y2: Y ex: x1 = x2 ey: y1 = y2
make_dirprod x1 y1 = make_dirprod x2 y2
X, Y: UU x1, x2: X y1, y2: Y ex: x1 = x2 ey: y1 = y2
make_dirprod x1 y1 = make_dirprod x2 y2
X, Y: UU x1: X y1, y2: Y ey: y1 = y2
make_dirprod x1 y1 = make_dirprod x1 y2
X, Y: UU x1: X y1: Y
make_dirprod x1 y1 = make_dirprod x1 y1
apply idpath.Defined.
A, B: UU ab, ab': A × B
pr1 ab = pr1 ab' → pr2 ab = pr2 ab' → ab = ab'
A, B: UU ab, ab': A × B
pr1 ab = pr1 ab' → pr2 ab = pr2 ab' → ab = ab'
A, B: UU ab, ab': A × B H: pr1 ab = pr1 ab' H': pr2 ab = pr2 ab'
ab = ab'
A, B: UU a: A b: B ab': A × B H: pr1 (a,, b) = pr1 ab' H': pr2 (a,, b) = pr2 ab'
a,, b = ab'
A, B: UU a: A b: B a': A b': B H: a = a' H': b = b'
a,, b = a',, b'
A, B: UU a: A b, b': B H': b = b'
a,, b = a,, b'
A, B: UU a: A b: B
a,, b = a,, b
apply idpath.Defined.(** *** The function [ maponpaths ] between paths types defined by a function between ambient typesand its behavior relative to [ @ ] and [ ! ] *)
T1, T2: UU f: T1 → T2 t1, t2: T1 e: t1 = t2
f t1 = f t2
T1, T2: UU f: T1 → T2 t1, t2: T1 e: t1 = t2
f t1 = f t2
T1, T2: UU f: T1 → T2 t1, t2: T1 e: t1 = t2
f t1 = f t2
T1, T2: UU f: T1 → T2 t1: T1
f t1 = f t1
apply idpath.Defined.(* useful with apply, to save typing *)
X, Y, Z: UU f: X → Y → Z x, x': X y, y': Y ex: x = x' ey: y = y'
f x y = f x' y'
X, Y, Z: UU f: X → Y → Z x, x': X y, y': Y ex: x = x' ey: y = y'
f x y = f x' y'
X, Y, Z: UU f: X → Y → Z x, x': X y, y': Y ex: x = x' ey: y = y'
f x y = f x' y'
X, Y, Z: UU f: X → Y → Z x: X y, y': Y ey: y = y'
f x y = f x y'
X, Y, Z: UU f: X → Y → Z x: X y: Y
f x y = f x y
apply idpath.Defined.
X, Y: UU x1, x2, x3: X f: X → Y e1: x1 = x2 e2: x2 = x3
maponpaths f (e1 @ e2) =
maponpaths f e1 @ maponpaths f e2
X, Y: UU x1, x2, x3: X f: X → Y e1: x1 = x2 e2: x2 = x3
maponpaths f (e1 @ e2) =
maponpaths f e1 @ maponpaths f e2
X, Y: UU x1, x2, x3: X f: X → Y e1: x1 = x2 e2: x2 = x3
maponpaths f (e1 @ e2) =
maponpaths f e1 @ maponpaths f e2
X, Y: UU x1, x3: X f: X → Y e2: x1 = x3
maponpaths f (idpath x1 @ e2) =
maponpaths f (idpath x1) @ maponpaths f e2
X, Y: UU x1: X f: X → Y
maponpaths f (idpath x1 @ idpath x1) =
maponpaths f (idpath x1) @ maponpaths f (idpath x1)
apply idpath.Defined.
X, Y: UU f: X → Y x1, x2: X e: x1 = x2
maponpaths f (! e) = ! maponpaths f e
X, Y: UU f: X → Y x1, x2: X e: x1 = x2
maponpaths f (! e) = ! maponpaths f e
X, Y: UU f: X → Y x1, x2: X e: x1 = x2
maponpaths f (! e) = ! maponpaths f e
X, Y: UU f: X → Y x1: X
maponpaths f (! idpath x1) =
! maponpaths f (idpath x1)
apply idpath.Defined.(** *** [ maponpaths ] for the identity functions and compositions of functions *)
X: UU x, x': X e: x = x'
maponpaths (idfun X) e = e
X: UU x, x': X e: x = x'
maponpaths (idfun X) e = e
X: UU x, x': X e: x = x'
maponpaths (idfun X) e = e
X: UU x: X
maponpaths (idfun X) (idpath x) = idpath x
apply idpath.Defined.
X, Y, Z: UU x, x': X f: X → Y g: Y → Z e: x = x'
maponpaths g (maponpaths f e) = maponpaths (g ∘ f) e
X, Y, Z: UU x, x': X f: X → Y g: Y → Z e: x = x'
maponpaths g (maponpaths f e) = maponpaths (g ∘ f) e
X, Y, Z: UU x, x': X f: X → Y g: Y → Z e: x = x'
maponpaths g (maponpaths f e) = maponpaths (g ∘ f) e
X, Y, Z: UU x: X f: X → Y g: Y → Z
maponpaths g (maponpaths f (idpath x)) =
maponpaths (g ∘ f) (idpath x)
apply idpath.Defined.(** *** Homotopy between sections *)Definitionhomot {X : UU} {P : X -> UU} (fg : ∏ x : X, P x) := ∏ x : X , f x = g x.Notation"f ~ g" := (homot f g).
X: UU P: X → UU f: ∏ x : X, P x
f ~ f
X: UU P: X → UU f: ∏ x : X, P x
f ~ f
X: UU P: X → UU f: ∏ x : X, P x x: X
f x = f x
apply idpath.Defined.Definitionhomotcomp {X:UU} {Y:X->UU} {ff'f'' : ∏ x : X, Y x}
(h : f ~ f') (h' : f' ~ f'') : f ~ f'' := λx, h x @ h' x.Definitioninvhomot {X:UU} {Y:X->UU} {ff' : ∏ x : X, Y x}
(h : f ~ f') : f' ~ f := λx, !(h x).Definitionfunhomot {XYZ:UU} (f : X -> Y) {gg' : Y -> Z}
(h : g ~ g') : g ∘ f ~ g' ∘ f := λx, h (f x).Definitionfunhomotsec {XY:UU} {Z:Y->UU} (f : X -> Y) {gg' : ∏ y:Y, Z y}
(h : g ~ g') : g ∘ f ~ g' ∘ f := λx, h (f x).Definitionhomotfun {XYZ : UU} {ff' : X -> Y} (h : f ~ f')
(g : Y -> Z) : g ∘ f ~ g ∘ f' := λx, maponpaths g (h x).(** *** Equality between functions defines a homotopy *)
T: UU P: T → UU f, g: ∏ t : T, P t
f = g → f ~ g
T: UU P: T → UU f, g: ∏ t : T, P t
f = g → f ~ g
T: UU P: T → UU f, g: ∏ t : T, P t h: f = g t: T
f t = g t
T: UU P: T → UU f: ∏ t : T, P t t: T
f t = f t
apply (idpath _).Defined.
T: UU P: T → UU f, g: ∏ t : T, P t
f = g → f ~ g
(* the same as toforallpaths, but with different implicit arguments *)
T: UU P: T → UU f, g: ∏ t : T, P t
f = g → f ~ g
T: UU P: T → UU f, g: ∏ t : T, P t e: f = g t: T
f t = g t
T: UU P: T → UU f: ∏ t : T, P t t: T
f t = f t
apply idpath.Defined.(** *** [ maponpaths ] for a function homotopic to the identityThe following three statements show that [ maponpaths ] defined bya function f which is homotopic to the identity is"surjective". It is later used to show that the maponpaths definedby a function which is a weak equivalence is itself a weakequivalence. *)(** Note that the type of the assumption h below can equivalently be written as[ homot f ( idfun X ) ] *)Definitionmaponpathshomidinv {X : UU} (f : X -> X)
(h : ∏ x : X, f x = x) (xx' : X) (e : f x = f x') :
x = x' := ! (h x) @ e @ (h x').
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: x = x'
maponpaths f e = h x @ e @ ! h x'
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: x = x'
maponpaths f e = h x @ e @ ! h x'
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: x = x'
maponpaths f e = h x @ e @ ! h x'
X: UU f: X → X h: ∏ x : X, f x = x x: X
maponpaths f (idpath x) = h x @ idpath x @ ! h x
X: UU f: X → X h: ∏ x : X, f x = x x: X
idpath (f x) = h x @ ! h x
X: UU f: X → X h: ∏ x : X, f x = x x: X
h x @ ! h x = idpath (f x)
apply pathsinv0r.Defined.
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x'
maponpaths f (maponpathshomidinv f h x x' e) = e
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x'
maponpaths f (maponpathshomidinv f h x x' e) = e
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x'
maponpaths f (maponpathshomidinv f h x x' e) = e
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x'
maponpaths f (! h x @ e @ h x') = e
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x'
h x @ (! h x @ e @ h x') @ ! h x' = e
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x'
∏ (X : UU) (a b c d : X) (p : a = b) (q : a = c)
(r : c = d), p @ (! p @ q @ r) @ ! r = q
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x' l: ∏ (X : UU) (a b c d : X) (p : a = b) (q : a = c)
(r : c = d), p @ (! p @ q @ r) @ ! r = q
h x @ (! h x @ e @ h x') @ ! h x' = e
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x'
∏ (X : UU) (a b c d : X) (p : a = b) (q : a = c)
(r : c = d), p @ (! p @ q @ r) @ ! r = q
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x' X0: UU a, b, c, d: X0 p: a = b q: a = c r: c = d
p @ (! p @ q @ r) @ ! r = q
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x' X0: UU a, c, d: X0 q: a = c r: c = d
idpath a @ (! idpath a @ q @ r) @ ! r = q
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x' X0: UU a, d: X0 r: a = d
idpath a @ (! idpath a @ idpath a @ r) @ ! r =
idpath a
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x' X0: UU a: X0
idpath a @
(! idpath a @ idpath a @ idpath a) @ ! idpath a =
idpath a
apply idpath.
X: UU f: X → X h: ∏ x : X, f x = x x, x': X e: f x = f x' l: ∏ (X : UU) (a b c d : X) (p : a = b) (q : a = c)
(r : c = d), p @ (! p @ q @ r) @ ! r = q
h x @ (! h x @ e @ h x') @ ! h x' = e
apply (l _ _ _ _ _ (h x) e (h x')).Defined.(** *** [ maponpaths ] in the case of a projection p with a section s *)(** Note that the type of the assumption eps below can equivalently be written as[ homot ( funcomp s p ) ( idfun X ) ] *)
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X y: Y e: s x = y
x = p y
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X y: Y e: s x = y
x = p y
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X y: Y e: s x = y
x = p y
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X y: Y e: s x = y
p (s x) = p y
apply (maponpaths p e).Defined.
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x, x': X e: s x = s x'
x = x'
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x, x': X e: s x = s x'
x = x'
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x, x': X e: s x = s x'
x = x'
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x, x': X e: s x = s x' e':= pathssec1 s p eps x (s x') e: x = p (s x')
x = x'
apply (e' @ (eps x')).Defined.
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
pathssec2 s p eps x x (idpath (s x)) = idpath x
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
pathssec2 s p eps x x (idpath (s x)) = idpath x
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
pathssec2 s p eps x x (idpath (s x)) = idpath x
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
pathssec1 s p eps x (s x) (idpath (s x)) @ eps x =
idpath x
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
(! eps x @ maponpaths p (idpath (s x))) @ eps x =
idpath x
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
(! eps x @ idpath (p (s x))) @ eps x = idpath x
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
∏ (X : UU) (a b : X) (p : a = b),
(! p @ idpath a) @ p = idpath b
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X e: ∏ (X : UU) (a b : X) (p : a = b),
(! p @ idpath a) @ p = idpath b
(! eps x @ idpath (p (s x))) @ eps x = idpath x
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
∏ (X : UU) (a b : X) (p : a = b),
(! p @ idpath a) @ p = idpath b
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X X0: UU a, b: X0 p0: a = b
(! p0 @ idpath a) @ p0 = idpath b
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X X0: UU a: X0
(! idpath a @ idpath a) @ idpath a = idpath a
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X X0: UU a: X0
idpath a = idpath a
apply idpath.
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X e: ∏ (X : UU) (a b : X) (p : a = b),
(! p @ idpath a) @ p = idpath b
(! eps x @ idpath (p (s x))) @ eps x = idpath x
apply e.Defined.
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x, x': X e: x = x'
pathssec2 s p eps x x' (maponpaths s e) = e
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x, x': X e: x = x'
pathssec2 s p eps x x' (maponpaths s e) = e
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x, x': X e: x = x'
pathssec2 s p eps x x' (maponpaths s e) = e
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
pathssec2 s p eps x x (maponpaths s (idpath x)) =
idpath x
X, Y: UU s: X → Y p: Y → X eps: ∏ x : X, p (s x) = x x: X
pathssec2 s p eps x x (idpath (s x)) = idpath x
apply pathssec2id.Defined.(** *** Fibrations and paths - the transport functions *)
X: UU P: X → UU x, x': X e: x = x'
∑ (f : P x → P x') (ee : ∏ p : P x, x,, p = x',, f p),
∏ pp : P x, maponpaths pr1 (ee pp) = e
X: UU P: X → UU x, x': X e: x = x'
∑ (f : P x → P x') (ee : ∏ p : P x, x,, p = x',, f p),
∏ pp : P x, maponpaths pr1 (ee pp) = e
X: UU P: X → UU x, x': X e: x = x'
∑ (f : P x → P x') (ee : ∏ p : P x, x,, p = x',, f p),
∏ pp : P x, maponpaths pr1 (ee pp) = e
X: UU P: X → UU x: X
∑ (f : P x → P x) (ee : ∏ p : P x, x,, p = x,, f p),
∏ pp : P x, maponpaths pr1 (ee pp) = idpath x
X: UU P: X → UU x: X
∑ ee : ∏ p : P x, x,, p = x,, idfun (P x) p,
∏ pp : P x, maponpaths pr1 (ee pp) = idpath x
X: UU P: X → UU x: X
∏ pp : P x,
maponpaths pr1 (idpath (x,, pp)) = idpath x
X: UU P: X → UU x: X
∏ pp : P x,
paths_rect (∑ y, P y) (x,, pp)
(λ (t2 : ∑ y, P y) (_ : x,, pp = t2),
pr1 (x,, pp) = pr1 t2) (idpath (pr1 (x,, pp)))
(x,, idfun (P x) pp) (idpath (x,, pp)) = idpath x
X: UU P: X → UU x: X
P x → idpath x = idpath x
X: UU P: X → UU x: X pp: P x
idpath x = idpath x
apply idpath.Defined.Definitiontransportf {X : UU} (P : X -> UU) {xx' : X}
(e : x = x') : P x -> P x' := pr1 (constr1 P e).Definitiontransportf_eq {X : UU} (P : X -> UU) {xx' : X} (e : x = x') ( p : P x ) :
tpair _ x p = tpair _ x' ( transportf P e p ) := ( pr1 ( pr2 ( constr1 P e ))) p .Definitiontransportb {X : UU} (P : X -> UU) {xx' : X}
(e : x = x') : P x' -> P x := transportf P (!e).Declare Scope transport.Notation"p # x" := (transportf _ p x) (only parsing) : transport.Notation"p #' x" := (transportb _ p x) (only parsing) : transport.Delimit Scope transport with transport.
X: UU P: X → UU x: X p: P x
transportf P (idpath x) p = p
X: UU P: X → UU x: X p: P x
transportf P (idpath x) p = p
X: UU P: X → UU x: X p: P x
transportf P (idpath x) p = p
apply idpath.Defined.
X, Y: UU f: X → Y P: Y → UU x, x': X e: x = x' p: P (f x)
transportf (λx : X, P (f x)) e p =
transportf P (maponpaths f e) p
X, Y: UU f: X → Y P: Y → UU x, x': X e: x = x' p: P (f x)
transportf (λx : X, P (f x)) e p =
transportf P (maponpaths f e) p
X, Y: UU f: X → Y P: Y → UU x, x': X e: x = x' p: P (f x)
transportf (λx : X, P (f x)) e p =
transportf P (maponpaths f e) p
X, Y: UU f: X → Y P: Y → UU x: X p: P (f x)
transportf (λx : X, P (f x)) (idpath x) p =
transportf P (maponpaths f (idpath x)) p
apply idpath.Defined.
X, Y: UU f: X → Y P: Y → UU x, x': X e: x' = x p: P (f x)
transportb (λx : X, P (f x)) e p =
transportb P (maponpaths f e) p
X, Y: UU f: X → Y P: Y → UU x, x': X e: x' = x p: P (f x)
transportb (λx : X, P (f x)) e p =
transportb P (maponpaths f e) p
X, Y: UU f: X → Y P: Y → UU x, x': X e: x' = x p: P (f x)
transportb (λx : X, P (f x)) e p =
transportb P (maponpaths f e) p
X, Y: UU f: X → Y P: Y → UU x': X p: P (f x')
transportb (λx : X, P (f x)) (idpath x') p =
transportb P (maponpaths f (idpath x')) p
apply idpath.Defined.
X: UU P: X → UU x, y, z: X e: y = x e': y = z p: P x
transportf P e' (transportb P e p) =
transportf P (! e @ e') p
X: UU P: X → UU x, y, z: X e: y = x e': y = z p: P x
transportf P e' (transportb P e p) =
transportf P (! e @ e') p
X: UU P: X → UU x, y, z: X e: y = x e': y = z p: P x
transportf P e' (transportb P e p) =
transportf P (! e @ e') p
X: UU P: X → UU x, y: X e: y = x p: P x
transportf P (idpath y) (transportb P e p) =
transportf P (! e @ idpath y) p
X: UU P: X → UU y: X p: P y
transportf P (idpath y) (transportb P (idpath y) p) =
transportf P (! idpath y @ idpath y) p
apply idpath.Defined.
X: UU P: X → UU x, y, z: X e: x = y e': z = y p: P x
transportb P e' (transportf P e p) =
transportf P (e @ ! e') p
X: UU P: X → UU x, y, z: X e: x = y e': z = y p: P x
transportb P e' (transportf P e p) =
transportf P (e @ ! e') p
X: UU P: X → UU x, y, z: X e: x = y e': z = y p: P x
transportb P e' (transportf P e p) =
transportf P (e @ ! e') p
X: UU P: X → UU x, z: X e: x = z p: P x
transportb P (idpath z) (transportf P e p) =
transportf P (e @ ! idpath z) p
X: UU P: X → UU x: X p: P x
transportb P (idpath x) (transportf P (idpath x) p) =
transportf P (idpath x @ ! idpath x) p
apply idpath.Defined.
X: UU P: X → UU x, y, z: X e: x = y e': y = z p: P x
transportf P e' (transportf P e p) =
transportf P (e @ e') p
X: UU P: X → UU x, y, z: X e: x = y e': y = z p: P x
transportf P e' (transportf P e p) =
transportf P (e @ e') p
X: UU P: X → UU x, y, z: X e: x = y e': y = z p: P x
transportf P e' (transportf P e p) =
transportf P (e @ e') p
X: UU P: X → UU x, y: X e: x = y p: P x
transportf P (idpath y) (transportf P e p) =
transportf P (e @ idpath y) p
X: UU P: X → UU x: X p: P x
transportf P (idpath x) (transportf P (idpath x) p) =
transportf P (idpath x @ idpath x) p
apply idpath.Defined.
X: UU P: X → UU x, y, z: X e: x = y e': y = z p: P z
transportb P e (transportb P e' p) =
transportb P (e @ e') p
X: UU P: X → UU x, y, z: X e: x = y e': y = z p: P z
transportb P e (transportb P e' p) =
transportb P (e @ e') p
X: UU P: X → UU x, y, z: X e: x = y e': y = z p: P z
transportb P e (transportb P e' p) =
transportb P (e @ e') p
X: UU P: X → UU x, y: X e: x = y p: P y
transportb P e (transportb P (idpath y) p) =
transportb P (e @ idpath y) p
X: UU P: X → UU x: X p: P x
transportb P (idpath x) (transportb P (idpath x) p) =
transportb P (idpath x @ idpath x) p
apply idpath.Defined.
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x x, y: X e: x = y p: P x
transportf Q e (f x p) = f y (transportf P e p)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x x, y: X e: x = y p: P x
transportf Q e (f x p) = f y (transportf P e p)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x x, y: X e: x = y p: P x
transportf Q e (f x p) = f y (transportf P e p)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x x: X p: P x
transportf Q (idpath x) (f x p) =
f x (transportf P (idpath x) p)
apply idpath.Defined.
X: UU P: X → UU f: ∏ x : X, P x x, y: X e: x = y
transportf P e (f x) = f y
X: UU P: X → UU f: ∏ x : X, P x x, y: X e: x = y
transportf P e (f x) = f y
X: UU P: X → UU f: ∏ x : X, P x x, y: X e: x = y
transportf P e (f x) = f y
exact (transport_map (P:= λ_,unit) (λx_,f x) e tt).Defined.
X, Y: UU P: X → UU x1, x2: X e: x1 = x2 f: P x1 → Y
transportf (λx : X, P x → Y) e f = f ∘ transportb P e
X, Y: UU P: X → UU x1, x2: X e: x1 = x2 f: P x1 → Y
transportf (λx : X, P x → Y) e f = f ∘ transportb P e
X, Y: UU P: X → UU x1, x2: X e: x1 = x2 f: P x1 → Y
transportf (λx : X, P x → Y) e f = f ∘ transportb P e
X, Y: UU P: X → UU x1: X f: P x1 → Y
transportf (λx : X, P x → Y) (idpath x1) f =
f ∘ transportb P (idpath x1)
apply idpath .Defined.
X: UU P: X → UU Z: UU x, x': X f: P x' → Z p: x = x' y: P x
f (transportf P p y) =
transportb (λx : X, P x → Z) p f y
X: UU P: X → UU Z: UU x, x': X f: P x' → Z p: x = x' y: P x
f (transportf P p y) =
transportb (λx : X, P x → Z) p f y
X: UU P: X → UU Z: UU x, x': X f: P x' → Z p: x = x' y: P x
f (transportf P p y) =
transportb (λx : X, P x → Z) p f y
X: UU P: X → UU Z: UU x: X f: P x → Z y: P x
f (transportf P (idpath x) y) =
transportb (λx : X, P x → Z) (idpath x) f y
apply idpath.Defined.
X: UU x1, x2: X e: x1 = x2 Y: UU
transportf (λ_ : X, Y) e = idfun Y
X: UU x1, x2: X e: x1 = x2 Y: UU
transportf (λ_ : X, Y) e = idfun Y
X: UU x1, x2: X e: x1 = x2 Y: UU
transportf (λ_ : X, Y) e = idfun Y
X: UU x1: X Y: UU
transportf (λ_ : X, Y) (idpath x1) = idfun Y
apply idpath.Defined.
X: UU x1, x2: X e: x1 = x2 Y: UU
transportb (λ_ : X, Y) e = idfun Y
X: UU x1, x2: X e: x1 = x2 Y: UU
transportb (λ_ : X, Y) e = idfun Y
X: UU x1, x2: X e: x1 = x2 Y: UU
transportb (λ_ : X, Y) e = idfun Y
X: UU x1: X Y: UU
transportb (λ_ : X, Y) (idpath x1) = idfun Y
apply idpath.Defined.
X: UU P: X → UU x1, x2: X e1, e2: x1 = x2 e: e1 = e2 p: P x1
transportf P e1 p = transportf P e2 p
X: UU P: X → UU x1, x2: X e1, e2: x1 = x2 e: e1 = e2 p: P x1
transportb P (idpath t) (transportf P (idpath t) p) =
p
apply idpath.Defined.
T: Type P: T → Type t, u: T e: t = u p: P u
transportf P e (transportb P e p) = p
T: Type P: T → Type t, u: T e: t = u p: P u
transportf P e (transportb P e p) = p
T: Type P: T → Type t, u: T e: t = u p: P u
transportf P e (transportb P e p) = p
T: Type P: T → Type t: T p: P t
transportf P (idpath t) (transportb P (idpath t) p) =
p
apply idpath.Defined.Close Scope transport.(** *** A series of lemmas about paths and [ total2 ] Some lemmas are adapted from the HoTT library http://github.com/HoTT/HoTT *)
A: UU B: A → UU a, b: ∑ y, B y
a = b → pr1 a = pr1 b
A: UU B: A → UU a, b: ∑ y, B y
a = b → pr1 a = pr1 b
A: UU B: A → UU a, b: ∑ y, B y X: a = b
pr1 a = pr1 b
apply maponpaths; assumption.Defined.
A, B, C: UU f: A → B → C a1: A b1: B a2: A b2: B p: a1 = a2 q: b1 = b2
f a1 b1 = f a2 b2
(* This lemma is an analogue of [maponpaths] for functions of two arguments. *)
A, B, C: UU f: A → B → C a1: A b1: B a2: A b2: B p: a1 = a2 q: b1 = b2
f a1 b1 = f a2 b2
A, B, C: UU f: A → B → C a1: A b1: B a2: A b2: B p: a1 = a2 q: b1 = b2
f a1 b1 = f a2 b2
A, B, C: UU f: A → B → C a1: A b1, b2: B q: b1 = b2
f a1 b1 = f a1 b2
A, B, C: UU f: A → B → C a1: A b1: B
f a1 b1 = f a1 b1
apply idpath.Defined.
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: transportf B p b1 = b2
f a1 b1 = f a2 b2
(* This lemma is a replacement for and a generalization of [total2_paths2_f], formerly called [total2_paths2], which does not refer to [total2]. The lemma [total2_paths2_f] can be obtained as the special case [f := tpair _], and Coq can often infer the value for [f], which is declared as an implicit argument. *)
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: transportf B p b1 = b2
f a1 b1 = f a2 b2
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: transportf B p b1 = b2
f a1 b1 = f a2 b2
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1, b2: B a1 q: transportf B (idpath a1) b1 = b2
f a1 b1 = f a1 b2
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1: B a1
f a1 b1 = f a1 (transportf B (idpath a1) b1)
apply idpath.Defined.
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: b1 = transportb B p b2
f a1 b1 = f a2 b2
(* This lemma is a replacement for and a generalization of [total2_paths2_b], which does not refer to [total2]. The lemma [total2_paths2_b] can be obtained as the special case [f := tpair _], and Coq can often infer the value for [f], which is declared as an implicit argument. *)
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: b1 = transportb B p b2
f a1 b1 = f a2 b2
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: b1 = transportb B p b2
f a1 b1 = f a2 b2
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1, b2: B a1 q: b1 = transportb B (idpath a1) b2
f a1 b1 = f a1 b2
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1, b2: B a1 q: b1 = b2
f a1 b1 = f a1 b2
A: UU B: A → UU C: UU f: ∏ a : A, B a → C a1: A b1: B a1
f a1 b1 = f a1 b1
apply idpath.Defined.
A, B: UU s, s': A × B p: pr1 s = pr1 s' q: pr2 s = pr2 s'
s = s'
A, B: UU s, s': A × B p: pr1 s = pr1 s' q: pr2 s = pr2 s'
s = s'
A, B: UU s, s': A × B p: pr1 s = pr1 s' q: pr2 s = pr2 s'
s = s'
A, B: UU a: A b: B a': A b': B p: a = a' q: b = b'
a,, b = a',, b'
exact (two_arg_paths p q).Defined.
A: UU B: A → UU s, s': ∑ x : A, B x p: pr1 s = pr1 s' q: transportf B p (pr2 s) = pr2 s'
s = s'
A: UU B: A → UU s, s': ∑ x : A, B x p: pr1 s = pr1 s' q: transportf B p (pr2 s) = pr2 s'
s = s'
A: UU B: A → UU s, s': ∑ x : A, B x p: pr1 s = pr1 s' q: transportf B p (pr2 s) = pr2 s'
s = s'
A: UU B: A → UU a: A b: B a a': A b': B a' p: a = a' q: transportf B p b = b'
a,, b = a',, b'
exact (two_arg_paths_f p q).Defined.
A: UU B: A → UU s, s': ∑ x : A, B x p: pr1 s = pr1 s' q: pr2 s = transportb B p (pr2 s')
s = s'
A: UU B: A → UU s, s': ∑ x : A, B x p: pr1 s = pr1 s' q: pr2 s = transportb B p (pr2 s')
s = s'
A: UU B: A → UU s, s': ∑ x : A, B x p: pr1 s = pr1 s' q: pr2 s = transportb B p (pr2 s')
s = s'
A: UU B: A → UU a: A b: B a a': A b': B a' p: a = a' q: b = transportb B p b'
a,, b = a',, b'
exact (two_arg_paths_b p q).Defined.
A, B: UU a1, a2: A b1, b2: B p: a1 = a2 q: b1 = b2
a1,, b1 = a2,, b2
A, B: UU a1, a2: A b1, b2: B p: a1 = a2 q: b1 = b2
a1,, b1 = a2,, b2
A, B: UU a1, a2: A b1, b2: B p: a1 = a2 q: b1 = b2
a1,, b1 = a2,, b2
exact (two_arg_paths p q).Defined.
A: UU B: A → UU a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: transportf B p b1 = b2
a1,, b1 = a2,, b2
A: UU B: A → UU a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: transportf B p b1 = b2
a1,, b1 = a2,, b2
A: UU B: A → UU a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: transportf B p b1 = b2
a1,, b1 = a2,, b2
exact (two_arg_paths_f p q).Defined.
A: UU B: A → UU a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: b1 = transportb B p b2
a1,, b1 = a2,, b2
A: UU B: A → UU a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: b1 = transportb B p b2
a1,, b1 = a2,, b2
A: UU B: A → UU a1: A b1: B a1 a2: A b2: B a2 p: a1 = a2 q: b1 = transportb B p b2
a1,, b1 = a2,, b2
exact (two_arg_paths_b p q).Defined.
X: UU P: X → UU x: X p, q: P x e: p = q
x,, p = x,, q
(* this function can often replaced by [maponpaths _] or by [maponpaths (tpair _ _)], except when the pairs in the goal have not been simplified enough to make the equality of their first parts evident, in which case this can be useful *)
X: UU P: X → UU x: X p, q: P x e: p = q
x,, p = x,, q
X: UU P: X → UU x: X p, q: P x e: p = q
x,, p = x,, q
X: UU P: X → UU x: X p, q: P x e: p = q
p = q
exact e.Defined.
A: UU B: A → UU u, v: ∑ x : A, B x p: u = v
transportf (λx : A, B x) (base_paths u v p) (pr2 u) =
pr2 v
A: UU B: A → UU u, v: ∑ x : A, B x p: u = v
transportf (λx : A, B x) (base_paths u v p) (pr2 u) =
pr2 v
A: UU B: A → UU u: ∑ x : A, B x
transportf (λx : A, B x) (base_paths u u (idpath u))
(pr2 u) = pr2 u
apply idpath.Defined.
A: UU B: A → UU x, y: ∑ x : A, B x p: x = y
total2_paths_f (base_paths x y p) (fiber_paths p) = p
A: UU B: A → UU x, y: ∑ x : A, B x p: x = y
total2_paths_f (base_paths x y p) (fiber_paths p) = p
A: UU B: A → UU x: ∑ x : A, B x
total2_paths_f (base_paths x x (idpath x))
(fiber_paths (idpath x)) = idpath x
base_paths (x,, H)
(x,, transportf (λx : A, B x) (idpath x) H)
(total2_paths_f (idpath x)
(idpath (transportf (λx : A, B x) (idpath x) H))) =
idpath x
apply idpath.Defined.
A: UU B: A → UU x, y: ∑ x : A, B x p: pr1 x = pr1 y q: transportf (λx : A, B x) p (pr2 x) = pr2 y
transportf
(λp' : pr1 x = pr1 y,
transportf (λx : A, B x) p' (pr2 x) = pr2 y)
(base_total2_paths q)
(fiber_paths (total2_paths_f p q)) = q
A: UU B: A → UU x, y: ∑ x : A, B x p: pr1 x = pr1 y q: transportf (λx : A, B x) p (pr2 x) = pr2 y
transportf
(λp' : pr1 x = pr1 y,
transportf (λx : A, B x) p' (pr2 x) = pr2 y)
(base_total2_paths q)
(fiber_paths (total2_paths_f p q)) = q
A: UU B: A → UU x: A H: B x y: ∑ x : A, B x p: pr1 (x,, H) = pr1 y q: transportf (λx : A, B x) p (pr2 (x,, H)) = pr2 y
transportf
(λp' : pr1 (x,, H) = pr1 y,
transportf (λx : A, B x) p' (pr2 (x,, H)) = pr2 y)
(base_total2_paths q)
(fiber_paths (total2_paths_f p q)) = q
A: UU B: A → UU x: A H: B x y: A K: B y p: pr1 (x,, H) = pr1 (y,, K) q: transportf (λx : A, B x) p (pr2 (x,, H)) =
pr2 (y,, K)
transportf
(λp' : pr1 (x,, H) = pr1 (y,, K),
transportf (λx : A, B x) p' (pr2 (x,, H)) =
pr2 (y,, K)) (base_total2_paths q)
(fiber_paths (total2_paths_f p q)) = q
A: UU B: A → UU x: A H: B x y: A K: B y p: x = y q: transportf (λx : A, B x) p H = K
transportf
(λp' : x = y, transportf (λx : A, B x) p' H = K)
(base_total2_paths q)
(fiber_paths (total2_paths_f p q)) = q
A: UU B: A → UU x: A H, K: B x q: transportf (λx : A, B x) (idpath x) H = K
transportf
(λp' : x = x, transportf (λx : A, B x) p' H = K)
(base_total2_paths q)
(fiber_paths (total2_paths_f (idpath x) q)) = q
A: UU B: A → UU x: A H: B x
transportf
(λp' : x = x,
transportf (λx : A, B x) p' H =
transportf (λx : A, B x) (idpath x) H)
(base_total2_paths
(idpath (transportf (λx : A, B x) (idpath x) H)))
(fiber_paths
(total2_paths_f (idpath x)
(idpath
(transportf (λx : A, B x) (idpath x) H)))) =
idpath (transportf (λx : A, B x) (idpath x) H)
apply idpath.Defined.
S, T: UU P: T → UU f: S → T
(∑ i : S, P (f i)) → ∑ j : T, P j
S, T: UU P: T → UU f: S → T
(∑ i : S, P (f i)) → ∑ j : T, P j
S, T: UU P: T → UU f: S → T x: ∑ i : S, P (f i)
∑ j : T, P j
exact (f(pr1 x),,pr2 x).Defined.
X: UU Y: X → UU a: X b: Y a e: ∏ x : X, Y x
a,, e a = a,, b → e a = b
(* this is called "Voldemort's theorem" by David McAllester, https://arxiv.org/pdf/1407.7274.pdf *)
X: UU Y: X → UU a: X b: Y a e: ∏ x : X, Y x
a,, e a = a,, b → e a = b
X: UU Y: X → UU a: X b: Y a e: ∏ x : X, Y x p: a,, e a = a,, b
e a = b
X: UU Y: X → UU a: X b: Y a e: ∏ x : X, Y x p: a,, e a = a,, b
e a =
transportf (λx : X, Y x)
(base_paths (a,, e a) (a,, b) p) (pr2 (a,, e a))
X: UU Y: X → UU a: X b: Y a e: ∏ x : X, Y x p: a,, e a = a,, b
e a =
transportf (λx : X, Y x) (maponpaths pr1 p)
(pr2 (a,, e a))
X: UU Y: X → UU a: X b: Y a e: ∏ x : X, Y x p: a,, e a = a,, b
e a =
transportf (λx : X, Y x) (maponpaths pr1 p) (e a)
apply pathsinv0, transport_section.Defined.(** *** Lemmas about transport adapted from the HoTT library and the HoTT book *)
A: UU B: A → UU C: ∏ a : A, B a → UU x1, x2: A p: x1 = x2 y: B x1 z: C x1 y
C x2 (transportf B p y)
A: UU B: A → UU C: ∏ a : A, B a → UU x1, x2: A p: x1 = x2 y: B x1 z: C x1 y
C x2 (transportf B p y)
A: UU B: A → UU C: ∏ a : A, B a → UU x1, x2: A p: x1 = x2 y: B x1 z: C x1 y
C x2 (transportf B p y)
A: UU B: A → UU C: ∏ a : A, B a → UU x1: A y: B x1 z: C x1 y
C x1 (transportf B (idpath x1) y)
exact z.Defined.
A: UU B: A → UU C: ∏ a : A, B a → UU x1, x2: A p: x1 = x2 yz: ∑ y : B x1, C x1 y
transportf (λx : A, ∑ y : B x, C x y) p yz =
transportf B p (pr1 yz),,
transportD B C p (pr1 yz) (pr2 yz)
A: UU B: A → UU C: ∏ a : A, B a → UU x1, x2: A p: x1 = x2 yz: ∑ y : B x1, C x1 y
transportf (λx : A, ∑ y : B x, C x y) p yz =
transportf B p (pr1 yz),,
transportD B C p (pr1 yz) (pr2 yz)
A: UU B: A → UU C: ∏ a : A, B a → UU x1, x2: A p: x1 = x2 yz: ∑ y : B x1, C x1 y
transportf (λx : A, ∑ y : B x, C x y) p yz =
transportf B p (pr1 yz),,
transportD B C p (pr1 yz) (pr2 yz)
A: UU B: A → UU C: ∏ a : A, B a → UU x1: A yz: ∑ y : B x1, C x1 y
transportf (λx : A, ∑ y : B x, C x y) (idpath x1) yz =
transportf B (idpath x1) (pr1 yz),,
transportD B C (idpath x1) (pr1 yz) (pr2 yz)
A: UU B: A → UU C: ∏ a : A, B a → UU x1: A pr1: B x1 pr2: C x1 pr1
transportf (λx : A, ∑ y : B x, C x y) (idpath x1)
(pr1,, pr2) =
transportf B (idpath x1) (Preamble.pr1 (pr1,, pr2)),,
transportD B C (idpath x1) (Preamble.pr1 (pr1,, pr2))
(Preamble.pr2 (pr1,, pr2))
apply idpath.Defined.
A: UU B, B': A → UU x, x': ∑ a : A, B a × B' a p: pr1 x = pr1 x'
transportf (λa : A, B a × B' a) p (pr2 x) =
make_dirprod
(transportf (λa : A, B a) p (pr1 (pr2 x)))
(transportf (λa : A, B' a) p (pr2 (pr2 x)))
A: UU B, B': A → UU x, x': ∑ a : A, B a × B' a p: pr1 x = pr1 x'
transportf (λa : A, B a × B' a) p (pr2 x) =
make_dirprod
(transportf (λa : A, B a) p (pr1 (pr2 x)))
(transportf (λa : A, B' a) p (pr2 (pr2 x)))
A: UU B, B': A → UU x, x': ∑ a : A, B a × B' a
transportf (λa : A, B a × B' a) (idpath (pr1 x))
(pr2 x) =
make_dirprod
(transportf (λa : A, B a) (idpath (pr1 x))
(pr1 (pr2 x)))
(transportf (λa : A, B' a) (idpath (pr1 x))
(pr2 (pr2 x)))
apply idpath.Defined.
A: UU B, B': A → UU x, x': ∑ a : A, B a × B' a p: pr1 x = pr1 x'
transportb (λa : A, B a × B' a) p (pr2 x') =
make_dirprod
(transportb (λa : A, B a) p (pr1 (pr2 x')))
(transportb (λa : A, B' a) p (pr2 (pr2 x')))
A: UU B, B': A → UU x, x': ∑ a : A, B a × B' a p: pr1 x = pr1 x'
transportb (λa : A, B a × B' a) p (pr2 x') =
make_dirprod
(transportb (λa : A, B a) p (pr1 (pr2 x')))
(transportb (λa : A, B' a) p (pr2 (pr2 x')))
A: UU B, B': A → UU x, x': ∑ a : A, B a × B' a p: pr1 x = pr1 x'
transportb (λa : A, B a × B' a) p (pr2 x') =
make_dirprod
(transportb (λa : A, B a) p (pr1 (pr2 x')))
(transportb (λa : A, B' a) p (pr2 (pr2 x')))
apply transportf_dirprod.Defined.
A: UU a, x1, x2: A p: x1 = x2 q: a = x1
transportf (λx : A, a = x) p q = q @ p
A: UU a, x1, x2: A p: x1 = x2 q: a = x1
transportf (λx : A, a = x) p q = q @ p
A: UU a, x1, x2: A p: x1 = x2 q: a = x1
transportf (λx : A, a = x) p q = q @ p
A: UU a, x1: A q: a = x1
transportf (λx : A, a = x) (idpath x1) q =
q @ idpath x1
A: UU a: A
transportf (λx : A, a = x) (idpath a) (idpath a) =
idpath a @ idpath a
apply idpath.Defined.
A: UU a, x1, x2: A p: x1 = x2 q: x1 = a
transportf (λx : A, x = a) p q = ! p @ q
A: UU a, x1, x2: A p: x1 = x2 q: x1 = a
transportf (λx : A, x = a) p q = ! p @ q
A: UU a, x1, x2: A p: x1 = x2 q: x1 = a
transportf (λx : A, x = a) p q = ! p @ q
A: UU a, x1: A q: x1 = a
transportf (λx : A, x = a) (idpath x1) q =
! idpath x1 @ q
A: UU x1: A
transportf (λx : A, x = x1) (idpath x1) (idpath x1) =
! idpath x1 @ idpath x1
apply idpath.Defined.
A: UU x1, x2: A p: x1 = x2 q: x1 = x1
transportf (λx : A, x = x) p q = ! p @ q @ p
A: UU x1, x2: A p: x1 = x2 q: x1 = x1
transportf (λx : A, x = x) p q = ! p @ q @ p
A: UU x1, x2: A p: x1 = x2 q: x1 = x1
transportf (λx : A, x = x) p q = ! p @ q @ p
A: UU x1: A q: x1 = x1
transportf (λx : A, x = x) (idpath x1) q =
! idpath x1 @ q @ idpath x1
A: UU x1: A q: x1 = x1
transportf (λx : A, x = x) (idpath x1) q =
q @ idpath x1
A: UU x1: A q: x1 = x1
q @ idpath x1 =
transportf (λx : A, x = x) (idpath x1) q
apply pathscomp0rid.Defined.(** *** Homotopies between families and the total spaces *)
X: UU P, Q: X → UU h: P ~ Q xp: ∑ y, P y
∑ y, Q y
X: UU P, Q: X → UU h: P ~ Q xp: ∑ y, P y
∑ y, Q y
X: UU P, Q: X → UU h: P ~ Q xp: ∑ y, P y
∑ y, Q y
X: UU P, Q: X → UU h: P ~ Q xp: ∑ y, P y
Q (pr1 xp)
X: UU P, Q: X → UU h: P ~ Q xp: ∑ y, P y
P (pr1 xp)
apply (pr2 xp).Defined.
X: UU P, Q: X → UU h1, h2: P ~ Q H: h1 ~ h2
famhomotfun h1 ~ famhomotfun h2
X: UU P, Q: X → UU h1, h2: P ~ Q H: h1 ~ h2
famhomotfun h1 ~ famhomotfun h2
X: UU P, Q: X → UU h1, h2: P ~ Q H: h1 ~ h2
famhomotfun h1 ~ famhomotfun h2
X: UU P, Q: X → UU h1, h2: P ~ Q H: h1 ~ h2 xp: ∑ y, P y
famhomotfun h1 xp = famhomotfun h2 xp
X: UU P, Q: X → UU h1, h2: P ~ Q H: h1 ~ h2 xp: ∑ y, P y
apply idpath.Defined.(** ** First fundamental notions *)(** *** Contractibility [ iscontr ] *)Definitioniscontr (T:UU) : UU := ∑ cntr:T, ∏ t:T, t=cntr.Notation"'∃!' x .. y , P"
:= (iscontr (∑ x, .. (∑ y, P) ..))
(at level200, x binder, y binder, right associativity) : type_scope.(* type this in emacs in agda-input method with \ex ! *)Definitionmake_iscontr {T : UU} : ∏ x : T, (∏ t : T, t = x) -> iscontr T
:= tpair _.Definitioniscontrpr1 {T : UU} : iscontr T -> T := pr1.Definitioniscontr_uniqueness {T} (i:iscontr T) (t:T) : t = iscontrpr1 i
:= pr2 i t.
X, Y: UU p: X → Y s: Y → X eps: ∏ y : Y, p (s y) = y is: iscontr X
iscontr Y
X, Y: UU p: X → Y s: Y → X eps: ∏ y : Y, p (s y) = y is: iscontr X
iscontr Y
X, Y: UU p: X → Y s: Y → X eps: ∏ y : Y, p (s y) = y is: iscontr X
iscontr Y
X, Y: UU p: X → Y s: Y → X eps: ∏ y : Y, p (s y) = y is: iscontr X x:= iscontrpr1 is: X
iscontr Y
X, Y: UU p: X → Y s: Y → X eps: ∏ y : Y, p (s y) = y is: iscontr X x:= iscontrpr1 is: X fe:= pr2 is: (λcntr : X, ∏ t : X, t = cntr) (pr1 is)
iscontr Y
X, Y: UU p: X → Y s: Y → X eps: ∏ y : Y, p (s y) = y is: iscontr X x:= iscontrpr1 is: X fe:= pr2 is: (λcntr : X, ∏ t : X, t = cntr) (pr1 is)
∏ t : Y, t = p x
X, Y: UU p: X → Y s: Y → X eps: ∏ y : Y, p (s y) = y is: iscontr X x:= iscontrpr1 is: X fe:= pr2 is: (λcntr : X, ∏ t : X, t = cntr) (pr1 is) t: Y
t = p x
apply (! (eps t) @ maponpaths p (fe (s t))).Defined.
X: UU is: iscontr X x, x': X
x = x'
X: UU is: iscontr X x, x': X
x = x'
X: UU is: iscontr X x, x': X
x = x'
apply ((pr2 is) x @ !(pr2 is x')).Defined.
A: UU B: A → UU isc: ∃! a : A, B a a: A p: B a
a = pr1 (pr1 isc)
A: UU B: A → UU isc: ∃! a : A, B a a: A p: B a
a = pr1 (pr1 isc)
A: UU B: A → UU isc: ∃! a : A, B a a: A p: B a Hi:= a,, p: ∑ y, B y
a = pr1 (pr1 isc)
apply (maponpaths pr1 (pr2 isc Hi)).Defined.(** *** Homotopy fibers [ hfiber ] *)Definitionhfiber {XY : UU} (f : X -> Y) (y : Y) : UU := ∑ x : X, f x = y.Definitionmake_hfiber {XY : UU} (f : X -> Y) {y : Y}
(x : X) (e : f x = y) : hfiber f y :=
tpair _ x e.Definitionhfiberpr1 {XY : UU} (f : X -> Y) (y : Y) : hfiber f y -> X := pr1.Definitionhfiberpr2 {XY : UU} (f : X -> Y) (y : Y) (y' : hfiber f y) : f (hfiberpr1 f y y') = y :=
pr2 y'.(** *** The functions between the hfibers of homotopic functions over the same point *)
X, Y: UU f, g: X → Y h: f ~ g y: Y
hfiber f y → hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y
hfiber f y → hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber f y
hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: f x = y
hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: f x = y
g x = y
apply (!(h x) @ e).Defined.
X, Y: UU f, g: X → Y h: f ~ g y: Y
hfiber g y → hfiber f y
X, Y: UU f, g: X → Y h: f ~ g y: Y
hfiber g y → hfiber f y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y
hfiber f y
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y
hfiber f y
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y
f x = y
apply (h x @ e).Defined.(** *** Paths in homotopy fibers *)
X, Y: UU f: X → Y y: Y xe1, xe2: hfiber f y e: xe1 = xe2
X, Y: UU f: X → Y y: Y xe1, xe2: hfiber f y ee: pr1 xe1 = pr1 xe2 eee: pr2 xe1 = maponpaths f ee @ pr2 xe2
xe1 = xe2
X, Y: UU f: X → Y y: Y xe1, xe2: hfiber f y ee: pr1 xe1 = pr1 xe2 eee: pr2 xe1 = maponpaths f ee @ pr2 xe2
xe1 = xe2
X, Y: UU f: X → Y y: Y xe1, xe2: hfiber f y ee: pr1 xe1 = pr1 xe2 eee: pr2 xe1 = maponpaths f ee @ pr2 xe2
xe1 = xe2
X, Y: UU f: X → Y y: Y t: X e1: f t = y xe2: hfiber f y ee: pr1 (t,, e1) = pr1 xe2 eee: pr2 (t,, e1) = maponpaths f ee @ pr2 xe2
t,, e1 = xe2
X, Y: UU f: X → Y y: Y t: X e1: f t = y t': X e2: f t' = y ee: pr1 (t,, e1) = pr1 (t',, e2) eee: pr2 (t,, e1) = maponpaths f ee @ pr2 (t',, e2)
t,, e1 = t',, e2
X, Y: UU f: X → Y y: Y t: X e1: f t = y t': X e2: f t' = y ee: t = t' eee: e1 = maponpaths f ee @ e2
t,, e1 = t',, e2
X, Y: UU f: X → Y y: Y t: X e1: f t = y t': X e2: f t' = y ee: t = t' eee: e1 = maponpaths f ee @ e2
make_hfiber f t e1 = t',, e2
X, Y: UU f: X → Y y: Y t: X e1: f t = y t': X e2: f t' = y ee: t = t' eee: e1 = maponpaths f ee @ e2
make_hfiber f t e1 = make_hfiber f t' e2
X, Y: UU f: X → Y y: Y t: X e1, e2: f t = y eee: e1 = maponpaths f (idpath t) @ e2
make_hfiber f t e1 = make_hfiber f t e2
X, Y: UU f: X → Y y: Y t: X e1, e2: f t = y eee: e1 = e2
make_hfiber f t e1 = make_hfiber f t e2
apply (maponpaths (λe, make_hfiber f t e) eee).Defined.(** *** Coconuses: spaces of paths that begin [ coconusfromt ] or end [ coconustot ] at a given point *)Definitioncoconusfromt (T : UU) (t : T) := ∑ t' : T, t = t'.Definitioncoconusfromtpair (T : UU) {tt' : T} (e: t = t') : coconusfromt T t
:= tpair _ t' e.Definitioncoconusfromtpr1 (T : UU) (t : T) : coconusfromt T t -> T := pr1.Definitioncoconustot (T : UU) (t : T) := ∑ t' : T, t' = t.Definitioncoconustotpair (T : UU) {tt' : T} (e: t' = t) : coconustot T t
:= tpair _ t' e.Definitioncoconustotpr1 (T : UU) (t : T) : coconustot T t -> T := pr1.(* There is a path between any two points in a coconus. As we always have a pointin a coconus, namely the one that is given by the pair of t and the path thatstarts at t and ends at t, the coconuses are contractible. *)
T: UU t: T c1, c2: coconustot T t
c1 = c2
T: UU t: T c1, c2: coconustot T t
c1 = c2
T: UU t: T c1, c2: coconustot T t
c1 = c2
T: UU t, x0: T x: x0 = t c2: coconustot T t
x0,, x = c2
T: UU x0: T c2: coconustot T x0
x0,, idpath x0 = c2
T: UU x0, x1: T y: x1 = x0
x0,, idpath x0 = x1,, y
T: UU x1: T
x1,, idpath x1 = x1,, idpath x1
apply idpath.Defined.
T: UU t: T
iscontr (coconustot T t)
T: UU t: T
iscontr (coconustot T t)
T: UU t: T
(λcntr : coconustot T t,
∏ t0 : coconustot T t, t0 = cntr) (t,, idpath t)
T: UU t, u: T
u,, idpath u = u,, idpath u
reflexivity.Defined.
T: UU t: T c1, c2: coconusfromt T t
c1 = c2
T: UU t: T c1, c2: coconusfromt T t
c1 = c2
T: UU t: T c1, c2: coconusfromt T t
c1 = c2
T: UU t, x0: T x: t = x0 c2: coconusfromt T t
x0,, x = c2
T: UU t: T c2: coconusfromt T t
t,, idpath t = c2
T: UU t, x1: T y: t = x1
t,, idpath t = x1,, y
T: UU t: T
t,, idpath t = t,, idpath t
apply idpath.Defined.
T: UU t: T
iscontr (coconusfromt T t)
T: UU t: T
iscontr (coconusfromt T t)
T: UU t: T
(λcntr : coconusfromt T t,
∏ t0 : coconusfromt T t, t0 = cntr) (t,, idpath t)
T: UU t, u: T
t,, idpath t = t,, idpath t
reflexivity.Defined.(** *** The total paths space of a type - two definitionsThe definitions differ by the (non) associativity of the [ total2 ]. *)Definitionpathsspace (T : UU) := ∑ t:T, coconusfromt T t.Definitionpathsspacetriple (T : UU) {t1t2 : T}
(e : t1 = t2) : pathsspace T := tpair _ t1 (coconusfromtpair T e).Definitiondeltap (T : UU) : T -> pathsspace T :=
λ (t : T), pathsspacetriple T (idpath t).Definitionpathsspace' (T : UU) := ∑ xy : T × T, pr1 xy = pr2 xy.(** *** Coconus of a function: the total space of the family of h-fibers *)Definitioncoconusf {XY : UU} (f : X -> Y) := ∑ y:Y, hfiber f y.Definitionfromcoconusf {XY : UU} (f : X -> Y) : coconusf f -> X :=
λyxe, pr1 (pr2 yxe).Definitiontococonusf {XY : UU} (f : X -> Y) : X -> coconusf f :=
λx, tpair _ (f x) (make_hfiber f x (idpath _)).
X, Y: UU f: X → Y
∏ yxe : coconusf f,
tococonusf f (fromcoconusf f yxe) = yxe
X, Y: UU f: X → Y
∏ yxe : coconusf f,
tococonusf f (fromcoconusf f yxe) = yxe
X, Y: UU f: X → Y yxe: coconusf f
tococonusf f (fromcoconusf f yxe) = yxe
X, Y: UU f: X → Y y: Y xe: hfiber f y
tococonusf f (fromcoconusf f (y,, xe)) = y,, xe
X, Y: UU f: X → Y y: Y x: X e: f x = y
tococonusf f (fromcoconusf f (y,, x,, e)) = y,, x,, e
X, Y: UU f: X → Y y: Y x: X e: f x = y
tococonusf f (pr1 (pr2 (y,, x,, e))) = y,, x,, e
X, Y: UU f: X → Y y: Y x: X e: f x = y
f (pr1 (pr2 (y,, x,, e))),,
make_hfiber f (pr1 (pr2 (y,, x,, e)))
(idpath (f (pr1 (pr2 (y,, x,, e))))) = y,, x,, e
X, Y: UU f: X → Y y: Y x: X e: f x = y
f x,, make_hfiber f x (idpath (f x)) = y,, x,, e
X, Y: UU f: X → Y x: X
f x,, make_hfiber f x (idpath (f x)) =
f x,, x,, idpath (f x)
apply idpath.Defined.
X, Y: UU f: X → Y
∏ x : X, fromcoconusf f (tococonusf f x) = x
X, Y: UU f: X → Y
∏ x : X, fromcoconusf f (tococonusf f x) = x
X, Y: UU f: X → Y x: X
fromcoconusf f (tococonusf f x) = x
X, Y: UU f: X → Y x: X
pr1 (pr2 (tococonusf f x)) = x
X, Y: UU f: X → Y x: X
pr1 (pr2 (f x,, make_hfiber f x (idpath (f x)))) = x
X, Y: UU f: X → Y x: X
x = x
apply idpath.Defined.(** ** Weak equivalences *)(** *** Basics - [ isweq ] and [ weq ] *)Definitionisweq {XY : UU} (f : X -> Y) : UU :=
∏ y : Y, iscontr (hfiber f y).
T: UU
isweq (idfun T)
T: UU
isweq (idfun T)
T: UU
isweq (idfun T)
T: UU
∏ y : T, iscontr (hfiber (idfun T) y)
T: UU y: T
iscontr (hfiber (idfun T) y)
T: UU y: T
∑ cntr : hfiber (idfun T) y,
∏ t : hfiber (idfun T) y, t = cntr
T: UU y: T
∏ t : hfiber (idfun T) y,
t = make_hfiber (idfun T) y (idpath y)
T: UU y: T t: hfiber (idfun T) y
t = make_hfiber (idfun T) y (idpath y)
T: UU y, x: T e: idfun T x = y
x,, e = make_hfiber (idfun T) y (idpath y)
T: UU x: T
x,, idpath (idfun T x) =
make_hfiber (idfun T) (idfun T x) (idpath (idfun T x))
apply idpath.Defined.Definitionweq (XY : UU) : UU := ∑ f:X->Y, isweq f.Notation"X ≃ Y" := (weq X%type Y%type) : type_scope.(* written \~- or \simeq in Agda input method *)Definitionpr1weq {XY : UU} := pr1 : X ≃ Y -> (X -> Y).Coercionpr1weq : weq >-> Funclass.Definitionweqproperty {XY} (f:X≃Y) : isweq f := pr2 f.
X, Y: UU w: X ≃ Y y: Y
hfiber w y
X, Y: UU w: X ≃ Y y: Y
hfiber w y
X, Y: UU w: X ≃ Y y: Y
hfiber w y
apply (iscontrpr1 (weqproperty w y)).Defined.
X, Y: UU w: X ≃ Y y: Y
∏ x : hfiber w y, x = weqccontrhfiber w y
X, Y: UU w: X ≃ Y y: Y
∏ x : hfiber w y, x = weqccontrhfiber w y
X, Y: UU w: X ≃ Y y: Y x: hfiber w y
x = weqccontrhfiber w y
X, Y: UU w: X ≃ Y y: Y x: hfiber w y
x = iscontrpr1 (weqproperty w y)
apply (pr2 (pr2 w y)).Defined.Definitionmake_weq {XY : UU} (f : X -> Y) (is: isweq f) : X ≃ Y :=
tpair (λf : X -> Y, isweq f) f is.Definitionidweq (X : UU) : X ≃ X :=
tpair (λf : X -> X, isweq f) (λ (x : X), x) (idisweq X).
X: UU f: X → ∅
isweq f
X: UU f: X → ∅
isweq f
X: UU f: X → ∅
isweq f
X: UU f: X → ∅ y: ∅
iscontr (hfiber f y)
apply (fromempty y).Defined.Definitionweqtoempty {X : UU} (f : X -> ∅) : X ≃ ∅ :=
make_weq _ (isweqtoempty f).
X, Y: UU f: X → Y is: ¬ Y
isweq f
X, Y: UU f: X → Y is: ¬ Y
isweq f
X, Y: UU f: X → Y is: ¬ Y
isweq f
X, Y: UU f: X → Y is: ¬ Y y: Y
iscontr (hfiber f y)
induction (is y).Defined.Definitionweqtoempty2 {XY : UU} (f : X -> Y) (is : ¬ Y) : X ≃ Y
:= make_weq _ (isweqtoempty2 f is).
X, Y: UU
¬ X → ¬ Y → X ≃ Y
X, Y: UU
¬ X → ¬ Y → X ≃ Y
X, Y: UU nx: ¬ X ny: ¬ Y
X ≃ Y
X, Y: UU nx: ¬ X ny: ¬ Y
X → Y
X, Y: UU nx: ¬ X ny: ¬ Y
isweq ?f
X, Y: UU nx: ¬ X ny: ¬ Y
X → Y
X, Y: UU nx: ¬ X ny: ¬ Y x: X
Y
X, Y: UU nx: ¬ X ny: ¬ Y x: X
X
exact x.
X, Y: UU nx: ¬ X ny: ¬ Y
isweq (λx : X, fromempty (nx x))
X, Y: UU nx: ¬ X ny: ¬ Y y: Y
iscontr (hfiber (λx : X, fromempty (nx x)) y)
X, Y: UU nx: ¬ X ny: ¬ Y y: Y
Y
exact y.Defined.Definitioninvmap {XY : UU} (w : X ≃ Y) : Y -> X :=
λ (y : Y), hfiberpr1 _ _ (weqccontrhfiber w y).(** *** Weak equivalences and paths spaces (more results in further sections) *)(** We now define different homotopies and maps between the paths spaces corresponding to a weak equivalence. What may look like unnecessary complexity in the definition of [ homotinvweqweq ] is due to the fact that the "naive" definition needs to be corrected in order for lemma [ homotweqinvweqweq ] to hold. *)
X, Y: UU w: X ≃ Y
∏ y : Y, w (invmap w y) = y
X, Y: UU w: X ≃ Y
∏ y : Y, w (invmap w y) = y
X, Y: UU w: X ≃ Y y: Y
w (invmap w y) = y
X, Y: UU w: X ≃ Y y: Y
w (hfiberpr1 w y (weqccontrhfiber w y)) = y
apply (pr2 (weqccontrhfiber w y)).Defined.
X, Y: UU w: X ≃ Y
∏ x : X, x = invmap w (w x)
X, Y: UU w: X ≃ Y
∏ x : X, x = invmap w (w x)
X, Y: UU w: X ≃ Y x: X
x = invmap w (w x)
X, Y: UU w: X ≃ Y x: X
x = hfiberpr1 w (w x) (weqccontrhfiber w (w x))
X, Y: UU w: X ≃ Y x: X xe1:= weqccontrhfiber w (w x): hfiber w (w x)
x = hfiberpr1 w (w x) xe1
X, Y: UU w: X ≃ Y x: X xe1:= weqccontrhfiber w (w x): hfiber w (w x) xe2:= make_hfiber w x (idpath (w x)): hfiber w (w x)
x = hfiberpr1 w (w x) xe1
X, Y: UU w: X ≃ Y x: X xe1:= weqccontrhfiber w (w x): hfiber w (w x) xe2:= make_hfiber w x (idpath (w x)): hfiber w (w x) p:= weqccontrhfiber2 w (w x) xe2: xe2 = weqccontrhfiber w (w x)
x = hfiberpr1 w (w x) xe1
apply (maponpaths pr1 p).Defined.Definitionhomotinvweqweq {XY : UU} (w : X ≃ Y) :
∏ x : X, invmap w (w x) = x
:= λ (x : X), ! (homotinvweqweq0 w x).Definitioninvmaponpathsweq {XY : UU} (w : X ≃ Y) (xx' : X) :
w x = w x' -> x = x'
:= pathssec2 w (invmap w) (homotinvweqweq w) x x'.Definitioninvmaponpathsweqid {XY : UU} (w : X ≃ Y) (x : X) :
invmaponpathsweq w _ _ (idpath (w x)) = idpath x
:= pathssec2id w (invmap w) (homotinvweqweq w) x.Definitionpathsweq1 {XY : UU} (w : X ≃ Y) (x : X) (y : Y) :
w x = y -> x = invmap w y
:= λe, maponpaths pr1 (pr2 (weqproperty w y) (x,,e)).Definitionpathsweq1' {XY : UU} (w : X ≃ Y) (x : X) (y : Y) :
x = invmap w y -> w x = y
:= λe, maponpaths w e @ homotweqinvweq w y.Definitionpathsweq3 {XY : UU} (w : X ≃ Y) {xx' : X}
(e : x = x') : invmaponpathsweq w x x' (maponpaths w e) = e
:= pathssec3 w (invmap w) (homotinvweqweq w) _.
X, Y: UU w: X ≃ Y x, x': X e: w x = w x'
maponpaths w (invmaponpathsweq w x x' e) = e
X, Y: UU w: X ≃ Y x, x': X e: w x = w x'
maponpaths w (invmaponpathsweq w x x' e) = e
X, Y: UU w: X ≃ Y x, x': X e: w x = w x'
maponpaths w (invmaponpathsweq w x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: (f,, is1) x = (f,, is1) x'
maponpaths (f,, is1)
(invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: (f,, is1) x = (f,, is1) x' w:= make_weq f is1: X ≃ Y
maponpaths (f,, is1)
(invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: (f,, is1) x = (f,, is1) x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X
maponpaths (f,, is1)
(invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: (f,, is1) x = (f,, is1) x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g ((f,, is1) x) = g ((f,, is1) x')
maponpaths (f,, is1)
(invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x')
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x'
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x'
maponpaths f eee = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x'
maponpaths f eee = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x'
maponpaths (g ∘ f) eee = ee
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee
maponpaths f eee = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x'
maponpaths (g ∘ f) eee = ee
apply maponpathshomid2.
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee
maponpaths f eee = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee
maponpaths g (maponpaths f eee) = maponpaths g e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee e3: maponpaths g (maponpaths f eee) = maponpaths g e
maponpaths f eee = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee
maponpaths g (maponpaths f eee) = maponpaths g e
apply (maponpathscomp f g eee @ e2).
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee e3: maponpaths g (maponpaths f eee) = maponpaths g e
maponpaths f eee = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee e3: maponpaths g (maponpaths f eee) = maponpaths g e s:= maponpaths g: f x = f x' → g (f x) = g (f x')
maponpaths f eee = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee e3: maponpaths g (maponpaths f eee) = maponpaths g e s:= maponpaths g: f x = f x' → g (f x) = g (f x') p:= pathssec2 g f (homotweqinvweq w) (f x) (f x'): g (f x) = g (f x') → f x = f x'
maponpaths f eee = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e2: maponpaths (g ∘ f) eee = ee e3: maponpaths g (maponpaths f eee) = maponpaths g e s:= maponpaths g: f x = f x' → g (f x) = g (f x') p:= pathssec2 g f (homotweqinvweq w) (f x) (f x'): g (f x) = g (f x') → f x = f x' eps:= pathssec3 g f (homotweqinvweq w): ∏ e : f x = f x',
pathssec2 g f (homotweqinvweq w) (f x) (f x') (maponpaths g e) = e
maponpaths f eee = e
apply (pathssec2 s p eps _ _ e3).
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e
maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e)
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e)
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e
maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e)
apply (maponpaths (λe0: f x = f x',
maponpaths f (invmaponpathsweq w x x' e0))
e1).
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e)
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e)
invmaponpathsweq w x x' (maponpaths f eee) = eee
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e) X0: invmaponpathsweq w x x' (maponpaths f eee) = eee
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e)
invmaponpathsweq w x x' (maponpaths f eee) = eee
apply (pathsweq3 w).
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e) X0: invmaponpathsweq w x x' (maponpaths f eee) = eee
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e) X0: invmaponpathsweq w x x' (maponpaths f eee) = eee
maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f eee
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e) X0: invmaponpathsweq w x x' (maponpaths f eee) = eee e5: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f eee
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e) X0: invmaponpathsweq w x x' (maponpaths f eee) = eee
maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f eee
apply (maponpaths (λeee0: x = x', maponpaths f eee0) X0).
X, Y: UU f: X → Y is1: isweq f x, x': X e: f x = f x' w:= make_weq f is1: X ≃ Y g:= invmap w: Y → X ee:= maponpaths g e: g (f x) = g (f x') eee:= maponpathshomidinv (g ∘ f) (homotinvweqweq w) x
x' ee: x = x' e1: maponpaths f eee = e e4: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f (invmaponpathsweq w x x' e) X0: invmaponpathsweq w x x' (maponpaths f eee) = eee e5: maponpaths f
(invmaponpathsweq w x x' (maponpaths f eee)) =
maponpaths f eee
maponpaths f (invmaponpathsweq (f,, is1) x x' e) = e
apply (! e4 @ e5 @ e1).Defined.
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z
f ~ g ∘ w → f ∘ invmap w ~ g
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z
f ~ g ∘ w → f ∘ invmap w ~ g
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z p: f ~ g ∘ w y: Y
(f ∘ invmap w) y = g y
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z y: Y
(g ∘ w) (invmap w y) = g y
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z y: Y
g (w (invmap w y)) = g y
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z y: Y
w (invmap w y) = y
apply homotweqinvweq.Defined.
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z
f ∘ invmap w ~ g → f ~ g ∘ w
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z
f ∘ invmap w ~ g → f ~ g ∘ w
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z q: f ∘ invmap w ~ g x: X
f x = (g ∘ w) x
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z q: f ∘ invmap w ~ g x: X
f x = (f ∘ invmap w) (w x)
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z q: f ∘ invmap w ~ g x: X
f x = f (invmap w (w x))
X: Type Y: UU Z: Type f: X → Z w: X ≃ Y g: Y → Z q: f ∘ invmap w ~ g x: X
invmap w (w x) = x
apply homotinvweqweq.Defined.
X, Y: UU v, w: X ≃ Y
invmap v ~ invmap w → v ~ w
X, Y: UU v, w: X ≃ Y
invmap v ~ invmap w → v ~ w
X, Y: UU v, w: X ≃ Y h: invmap v ~ invmap w x: X
v x = w x
X, Y: UU v, w: X ≃ Y h: invmap v ~ invmap w x: X
v x = w (invmap w (v x))
X, Y: UU v, w: X ≃ Y h: invmap v ~ invmap w x: X
w (invmap w (v x)) = w x
X, Y: UU v, w: X ≃ Y h: invmap v ~ invmap w x: X
v x = w (invmap w (v x))
X, Y: UU v, w: X ≃ Y h: invmap v ~ invmap w x: X
w (invmap w (v x)) = v x
apply homotweqinvweq.
X, Y: UU v, w: X ≃ Y h: invmap v ~ invmap w x: X
w (invmap w (v x)) = w x
X, Y: UU v, w: X ≃ Y h: invmap v ~ invmap w x: X
w (invmap v (v x)) = w x
X, Y: UU v, w: X ≃ Y h: invmap v ~ invmap w x: X
w x = w x
apply idpath.Defined.
X, Y: Type v, w: X → Y v', w': Y → X
w ∘ w' ~ idfun Y → v' ∘ v ~ idfun X → v' ~ w' → v ~ w
X, Y: Type v, w: X → Y v', w': Y → X
w ∘ w' ~ idfun Y → v' ∘ v ~ idfun X → v' ~ w' → v ~ w
X, Y: Type v, w: X → Y v', w': Y → X p: w ∘ w' ~ idfun Y q: v' ∘ v ~ idfun X h: v' ~ w' x: X
v x = w x
X, Y: Type v, w: X → Y v', w': Y → X p: w ∘ w' ~ idfun Y q: v' ∘ v ~ idfun X h: v' ~ w' x: X
v x = w (w' (v x))
X, Y: Type v, w: X → Y v', w': Y → X p: w ∘ w' ~ idfun Y q: v' ∘ v ~ idfun X h: v' ~ w' x: X
w (w' (v x)) = w x
X, Y: Type v, w: X → Y v', w': Y → X p: w ∘ w' ~ idfun Y q: v' ∘ v ~ idfun X h: v' ~ w' x: X
v x = w (w' (v x))
X, Y: Type v, w: X → Y v', w': Y → X p: w ∘ w' ~ idfun Y q: v' ∘ v ~ idfun X h: v' ~ w' x: X
w (w' (v x)) = v x
apply p.
X, Y: Type v, w: X → Y v', w': Y → X p: w ∘ w' ~ idfun Y q: v' ∘ v ~ idfun X h: v' ~ w' x: X
w (w' (v x)) = w x
X, Y: Type v, w: X → Y v', w': Y → X p: w ∘ w' ~ idfun Y q: v' ∘ v ~ idfun X h: v' ~ w' x: X
w' (v x) = x
X, Y: Type v, w: X → Y v', w': Y → X p: w ∘ w' ~ idfun Y q: v' ∘ v ~ idfun X h: v' ~ w' x: X
v' (v x) = x
apply q.Defined.(** *** Adjointness property of a weak equivalence and its inverse *)
X, Y: UU f: X → Y x, x': X e1: x = x' e2: f x' = f x ee: idpath (f x) = maponpaths f e1 @ e2
maponpaths f (! e1) = e2
X, Y: UU f: X → Y x, x': X e1: x = x' e2: f x' = f x ee: idpath (f x) = maponpaths f e1 @ e2
maponpaths f (! e1) = e2
X, Y: UU f: X → Y x, x': X e1: x = x' e2: f x' = f x ee: idpath (f x) = maponpaths f e1 @ e2
maponpaths f (! e1) = e2
X, Y: UU f: X → Y x: X e2: f x = f x ee: idpath (f x) = maponpaths f (idpath x) @ e2
maponpaths f (! idpath x) = e2
X, Y: UU f: X → Y x: X e2: f x = f x ee: idpath (f x) = e2
idpath (f x) = e2
exact ee.Defined.(* this is the adjointness relation for w and its homotopy inverse: *)
X, Y: UU w: X ≃ Y x: X
maponpaths w (homotinvweqweq w x) =
homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y x: X
maponpaths w (homotinvweqweq w x) =
homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y x: X
maponpaths w (homotinvweqweq w x) =
homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y x: X
maponpaths w (! homotinvweqweq0 w x) =
homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y x: X
maponpaths w
(! maponpaths pr1
(weqccontrhfiber2 w (w x)
(make_hfiber w x (idpath (w x))))) =
homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y x: X hfid:= make_hfiber w x (idpath (w x)): hfiber w (w x)
maponpaths w
(! maponpaths pr1 (weqccontrhfiber2 w (w x) hfid)) =
homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y x: X hfid:= make_hfiber w x (idpath (w x)): hfiber w (w x) hfcc:= weqccontrhfiber w (w x): hfiber w (w x)
maponpaths w
(! maponpaths pr1 (weqccontrhfiber2 w (w x) hfid)) =
homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y x: X hfid:= make_hfiber w x (idpath (w x)): hfiber w (w x) hfcc:= weqccontrhfiber w (w x): hfiber w (w x)
maponpaths w
(! maponpaths pr1 (weqccontrhfiber2 w (w x) hfid)) =
pr2 (weqccontrhfiber w (w x))
X, Y: UU w: X ≃ Y x: X hfid:= make_hfiber w x (idpath (w x)): hfiber w (w x) hfcc:= weqccontrhfiber w (w x): hfiber w (w x)
idpath (w x) =
maponpaths w
(maponpaths pr1 (weqccontrhfiber2 w (w x) hfid)) @
pr2 (weqccontrhfiber w (w x))
apply (@hfibertriangle1 _ _ w _ hfid hfcc (weqccontrhfiber2 w (w x) hfid)).Defined.(* another way the adjointness relation may occur (added by D. Grayson, Oct. 2015): *)
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
transportf (P ∘ w) (! homotinvweqweq w x) p =
transportf P (! homotweqinvweq w (w x)) p
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
transportf (P ∘ w) (! homotinvweqweq w x) p =
transportf P (! homotweqinvweq w (w x)) p
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
transportf (P ∘ w) (! homotinvweqweq w x) p =
transportf P (! homotweqinvweq w (w x)) p
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
transportf P (maponpaths w (! homotinvweqweq w x)) p =
transportf P (! homotweqinvweq w (w x)) p
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
maponpaths w (! homotinvweqweq w x) =
! homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
! maponpaths w (homotinvweqweq w x) =
! homotweqinvweq w (w x)
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
maponpaths w (homotinvweqweq w x) =
homotweqinvweq w (w x)
apply homotweqinvweqweq.Defined.
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
transportb (P ∘ w) (homotinvweqweq w x) p =
transportb P (homotweqinvweq w (w x)) p
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
transportb (P ∘ w) (homotinvweqweq w x) p =
transportb P (homotweqinvweq w (w x)) p
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
transportb (P ∘ w) (homotinvweqweq w x) p =
transportb P (homotweqinvweq w (w x)) p
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
transportb P (maponpaths w (homotinvweqweq w x)) p =
transportb P (homotweqinvweq w (w x)) p
X, Y: UU w: X ≃ Y P: Y → UU x: X p: P (w x)
maponpaths w (homotinvweqweq w x) =
homotweqinvweq w (w x)
apply homotweqinvweqweq.Defined.(** *** Transport functions are weak equivalences *)
apply (iscontrretract (invmap w) w (homotinvweqweq w) is).Defined.(** *** [ unit ] and contractibility *)(** [ unit ] is contractible (recall that [ tt ] is the name of the canonical term of the type [ unit ]). *)
apply (isProofIrrelevantUnit t tt).Defined.(** [ paths ] in [ unit ] are contractible. *)
x, x': unit
iscontr (x = x')
x, x': unit
iscontr (x = x')
x, x': unit
iscontr (x = x')
x, x': unit
∏ t : x = x', t = isProofIrrelevantUnit x x'
x, x': unit e': x = x'
e' = isProofIrrelevantUnit x x'
x': unit e': tt = x'
e' = isProofIrrelevantUnit tt x'
e': tt = tt
e' = isProofIrrelevantUnit tt tt
e': tt = tt
e' = idpath tt
apply unitl3.Defined.(** A type [ T : UU ] is contractible if and only if [ T -> unit ] is a weak equivalence. *)
e1, e2: tt = tt
e1 = e2
e1, e2: tt = tt
e1 = e2
e1, e2: tt = tt
e1 = e2
e1, e2: tt = tt
iscontr (tt = tt)
apply (iscontrpathsinunit tt tt).Defined.
T: UU is: iscontr T
isweq (λ_ : T, tt)
T: UU is: iscontr T
isweq (λ_ : T, tt)
T: UU is: iscontr T
isweq (λ_ : T, tt)
T: UU is: iscontr T
∏ y : unit, iscontr (hfiber (λ_ : T, tt) y)
T: UU is: iscontr T y: unit
iscontr (hfiber (λ_ : T, tt) y)
T: UU is: iscontr T
iscontr (hfiber (λ_ : T, tt) tt)
T: UU c: T h: ∏ t : T, t = c
iscontr (hfiber (λ_ : T, tt) tt)
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt
iscontr (hfiber (λ_ : T, tt) tt)
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt
∏ t : hfiber (λ_ : T, tt) tt, t = hc
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt ha: hfiber (λ_ : T, tt) tt
ha = hc
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt x: T e: tt = tt
x,, e = hc
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt x: T e: tt = tt
x,, e = make_hfiber (λ_ : T, tt) c (idpath tt)
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt x: T e: tt = tt
x,, e = c,, idpath tt
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt x: T e: tt = tt
x,, e = c,, idpath tt
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt x: T e: tt = tt
x,, e = c,, idpath tt
T: UU c: T h: ∏ t : T, t = c hc:= make_hfiber (λ_ : T, tt) c (idpath tt): hfiber (λ_ : T, tt) tt x: T e: tt = tt
transportf (λ_ : T, tt = tt) (h x) e = idpath tt
apply ifcontrthenunitl0.Defined.Definitionweqcontrtounit {T : UU} (is : iscontr T) : T ≃ unit :=
make_weq _ (isweqcontrtounit is).
X: UU w: X ≃ unit
iscontr X
X: UU w: X ≃ unit
iscontr X
X: UU w: X ≃ unit
iscontr X
X: UU w: X ≃ unit
iscontr unit
apply iscontrunit.Defined.(** *** Homotopy equivalence is a weak equivalence *)Definitionhfibersgftog {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(xe : hfiber (g ∘ f) z) : hfiber g z :=
make_hfiber g (f (pr1 xe)) (pr2 xe).
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X xe: hfiber g x0
∑ xe' : hfiber (g ∘ f) x0,
xe = hfibersgftog f g x0 xe'
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X xe: hfiber g x0
∑ xe' : hfiber (g ∘ f) x0,
xe = hfibersgftog f g x0 xe'
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X xe: hfiber g x0
∑ xe' : hfiber (g ∘ f) x0,
xe = hfibersgftog f g x0 xe'
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0
∑ xe' : hfiber (g ∘ f) x0,
y0,, e = hfibersgftog f g x0 xe'
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0 eint:= pathssec1 g f efg y0 x0 e: y0 = f x0
∑ xe' : hfiber (g ∘ f) x0,
y0,, e = hfibersgftog f g x0 xe'
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0 eint:= pathssec1 g f efg y0 x0 e: y0 = f x0 ee:= ! maponpaths g eint @ e: g (f x0) = x0
∑ xe' : hfiber (g ∘ f) x0,
y0,, e = hfibersgftog f g x0 xe'
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0 eint:= pathssec1 g f efg y0 x0 e: y0 = f x0 ee:= ! maponpaths g eint @ e: g (f x0) = x0
y0,, e =
hfibersgftog f g x0 (make_hfiber (g ∘ f) x0 ee)
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0 eint:= pathssec1 g f efg y0 x0 e: y0 = f x0 ee:= ! maponpaths g eint @ e: g (f x0) = x0
y0,, e =
make_hfiber g (f (pr1 (make_hfiber (g ∘ f) x0 ee)))
(pr2 (make_hfiber (g ∘ f) x0 ee))
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0 eint:= pathssec1 g f efg y0 x0 e: y0 = f x0 ee:= ! maponpaths g eint @ e: g (f x0) = x0
y0,, e = f (pr1 (x0,, ee)),, pr2 (x0,, ee)
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0 eint:= pathssec1 g f efg y0 x0 e: y0 = f x0 ee:= ! maponpaths g eint @ e: g (f x0) = x0
y0,, e = f x0,, ee
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0 eint:= pathssec1 g f efg y0 x0 e: y0 = f x0 ee:= ! maponpaths g eint @ e: g (f x0) = x0
transportf (λx : Y, g x = x0) eint e = ee
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X y0: Y e: g y0 = x0 ee:= ! maponpaths g (idpath y0) @ e: g y0 = x0
transportf (λx : Y, g x = x0) (idpath y0) e = ee
apply idpath.Defined.
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X is: iscontr (hfiber (g ∘ f) x0)
iscontr (hfiber g x0)
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X is: iscontr (hfiber (g ∘ f) x0)
iscontr (hfiber g x0)
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X is: iscontr (hfiber (g ∘ f) x0)
iscontr (hfiber g x0)
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X is: iscontr (hfiber (g ∘ f) x0) f1:= hfibersgftog f g x0: hfiber (g ∘ f) x0 → hfiber g x0
iscontr (hfiber g x0)
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X is: iscontr (hfiber (g ∘ f) x0) f1:= hfibersgftog f g x0: hfiber (g ∘ f) x0 → hfiber g x0 g1:= λxe : hfiber g x0, pr1 (constr2 f g efg x0 xe): hfiber g x0 → hfiber (g ∘ f) x0
iscontr (hfiber g x0)
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X is: iscontr (hfiber (g ∘ f) x0) f1:= hfibersgftog f g x0: hfiber (g ∘ f) x0 → hfiber g x0 g1:= λxe : hfiber g x0, pr1 (constr2 f g efg x0 xe): hfiber g x0 → hfiber (g ∘ f) x0 efg1:= λxe : hfiber g x0, ! pr2 (constr2 f g efg x0 xe): ∏ xe : hfiber g x0,
hfibersgftog f g x0
(pr1 (constr2 f g efg x0 xe)) = xe
iscontr (hfiber g x0)
X, Y: UU f: X → Y g: Y → X efg: ∏ y : Y, f (g y) = y x0: X is: iscontr (hfiber (g ∘ f) x0) f1:= hfibersgftog f g x0: hfiber (g ∘ f) x0 → hfiber g x0 g1:= λxe : hfiber g x0, pr1 (constr2 f g efg x0 xe): hfiber g x0 → hfiber (g ∘ f) x0 efg1:= λxe : hfiber g x0, ! pr2 (constr2 f g efg x0 xe): ∏ xe : hfiber g x0,
hfibersgftog f g x0
(pr1 (constr2 f g efg x0 xe)) = xe
iscontr (hfiber (g ∘ f) x0)
applyis.Defined.
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber f y
hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber f y
hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber f y
hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber f y x:= pr1 xe: X
hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber f y x:= pr1 xe: X e:= pr2 xe: (λx : X, f x = y) (pr1 xe)
hfiber g y
apply (make_hfiber g x (!(h x) @ e)).Defined.
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y
hfiber f y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y
hfiber f y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y
hfiber f y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y x:= pr1 xe: X
hfiber f y
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y x:= pr1 xe: X e:= pr2 xe: (λx : X, g x = y) (pr1 xe)
hfiber f y
apply (make_hfiber f x (h x @ e)).Defined.
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y
homothfiber1 f g h y (homothfiber2 f g h y xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y
homothfiber1 f g h y (homothfiber2 f g h y xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y xe: hfiber g y
homothfiber1 f g h y (homothfiber2 f g h y xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y
homothfiber1 f g h y (homothfiber2 f g h y (x,, e)) =
x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y
homothfiber1 f g h y (homothfiber2 f g h y (x,, e)) =
x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y
homothfiber1 f g h y
(homothfiber2 f g h y (make_hfiber g x e)) =
make_hfiber g x e
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y
homothfiber1 f g h y
(homothfiber2 f g h y (make_hfiber g x e)) =
make_hfiber g x e
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y
homothfiber1 f g h y (homothfiber2 f g h y xe2) = xe2
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y
! h x @ h x @ e = e
(* A little lemma: *)
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y
∏ (a b c : Y) (p : a = b) (q : b = c), ! p @ p @ q = q
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y ee: ∏ (a b c : Y) (p : a = b) (q : b = c),
! p @ p @ q = q
! h x @ h x @ e = e
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y
∏ (a b c : Y) (p : a = b) (q : b = c), ! p @ p @ q = q
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y a, b, c: Y p: a = b q: b = c
! p @ p @ q = q
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y a, c: Y q: a = c
! idpath a @ idpath a @ q = q
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y a: Y
! idpath a @ idpath a @ idpath a = idpath a
apply idpath.
X, Y: UU f, g: X → Y h: f ~ g y: Y x: X e: g x = y xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y ee: ∏ (a b c : Y) (p : a = b) (q : b = c),
! p @ p @ q = q
! h x @ h x @ e = e
apply ee.Defined.
X, Y: UU f, g: X → Y h: f ~ g y: Y is: iscontr (hfiber f y)
iscontr (hfiber g y)
X, Y: UU f, g: X → Y h: f ~ g y: Y is: iscontr (hfiber f y)
iscontr (hfiber g y)
X, Y: UU f, g: X → Y h: f ~ g y: Y is: iscontr (hfiber f y)
iscontr (hfiber g y)
X, Y: UU f, g: X → Y h: f ~ g y: Y is: iscontr (hfiber f y) a:= homothfiber1 f g h y: hfiber f y → hfiber g y
iscontr (hfiber g y)
X, Y: UU f, g: X → Y h: f ~ g y: Y is: iscontr (hfiber f y) a:= homothfiber1 f g h y: hfiber f y → hfiber g y b:= homothfiber2 f g h y: hfiber g y → hfiber f y
iscontr (hfiber g y)
X, Y: UU f, g: X → Y h: f ~ g y: Y is: iscontr (hfiber f y) a:= homothfiber1 f g h y: hfiber f y → hfiber g y b:= homothfiber2 f g h y: hfiber g y → hfiber f y eab:= homothfiberretr f g h y: ∏ xe : hfiber g y, homothfiber1 f g h y (homothfiber2 f g h y xe) = xe
iscontr (hfiber g y)
apply (iscontrretract a b eab is).Defined.
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f1 → isweq f2
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f1 → isweq f2
X, Y: UU f1, f2: X → Y h: f1 ~ f2 x0: isweq f1
isweq f2
X, Y: UU f1, f2: X → Y h: f1 ~ f2 x0: isweq f1
∏ y : Y, iscontr (hfiber f2 y)
X, Y: UU f1, f2: X → Y h: f1 ~ f2 x0: isweq f1 y: Y
iscontr (hfiber f2 y)
X, Y: UU f1, f2: X → Y h: f1 ~ f2 x0: isweq f1 y: Y
iscontr (hfiber f1 y)
apply x0.Defined.
X, Y: UU f: X ≃ Y g: X → Y
f ~ g → X ≃ Y
(* this lemma may be used to replace an equivalence by one whose forward map has a simpler definition, keeping the same inverse map, judgmentally *)
X, Y: UU f: X ≃ Y g: X → Y
f ~ g → X ≃ Y
X, Y: UU f: X ≃ Y g: X → Y e: f ~ g
X ≃ Y
exact (g ,, isweqhomot f g e (weqproperty f)).Defined.
X, Y: UU f1: X ≃ Y f2: X → Y e: f1 ~ f2
remakeweq e = f2
(* check the claim in the comment above *)
X, Y: UU f1: X ≃ Y f2: X → Y e: f1 ~ f2
remakeweq e = f2
X, Y: UU f1: X ≃ Y f2: X → Y e: f1 ~ f2
remakeweq e = f2
apply idpath.Defined.
X, Y: UU f1: X ≃ Y f2: X → Y e: f1 ~ f2
invmap (remakeweq e) = invmap f1
(* check the claim in the comment above *)
X, Y: UU f1: X ≃ Y f2: X → Y e: f1 ~ f2
invmap (remakeweq e) = invmap f1
X, Y: UU f1: X ≃ Y f2: X → Y e: f1 ~ f2
invmap (remakeweq e) = invmap f1
apply idpath.Defined.
X: Type
X → iscontr X → iscontr X
X: Type
X → iscontr X → iscontr X
X: Type x: X i: iscontr X
iscontr X
X: Type x: X i: iscontr X
∏ t : X, t = x
X: Type x: X i: iscontr X y: X
y = x
X: Type x: X i: iscontr X y: X
iscontr X
exact i.Defined.
X: Type x: X i: iscontr X
iscontrpr1 (iscontr_move_point x i) = x
X: Type x: X i: iscontr X
iscontrpr1 (iscontr_move_point x i) = x
X: Type x: X i: iscontr X
iscontrpr1 (iscontr_move_point x i) = x
apply idpath.Defined.
X, Y: UU f: X ≃ Y h: Y → X
invmap f ~ h → X ≃ Y
(* this lemma may be used to replace an equivalence by one whose inverse map is simpler, leaving the forward map the same, judgmentally *)
X, Y: UU f: X ≃ Y h: Y → X
invmap f ~ h → X ≃ Y
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h
X ≃ Y
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h
isweq f
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h y: Y
iscontr (hfiber f y)
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h y: Y
hfiber f y
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h y: Y p: hfiber f y
iscontr (hfiber f y)
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h y: Y
hfiber f y
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h y: Y
f (h y) = y
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h y: Y
invmap f y = h y
apply e.
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h y: Y p: hfiber f y
iscontr (hfiber f y)
exact (iscontr_move_point p (weqproperty f y)).Defined.
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h
remakeweqinv e = f
(* check the claim in the comment above *)
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h
remakeweqinv e = f
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h
remakeweqinv e = f
apply idpath.Defined.
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h
invmap (remakeweqinv e) = h
(* check the claim in the comment above *)
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h
invmap (remakeweqinv e) = h
X, Y: UU f: X ≃ Y h: Y → X e: invmap f ~ h
invmap (remakeweqinv e) = h
apply idpath.Defined.
X, Y: UU f: X ≃ Y g: X → Y h: Y → X
f ~ g → invmap f ~ h → X ≃ Y
(* this lemma may be used to replace an equivalence by one whose two maps are simpler *)
X, Y: UU f: X ≃ Y g: X → Y h: Y → X
f ~ g → invmap f ~ h → X ≃ Y
X, Y: UU f: X ≃ Y g: X → Y h: Y → X r: f ~ g s: invmap f ~ h
X ≃ Y
use (remakeweqinv (f := remakeweq r) s).Defined.
X, Y: UU f: X ≃ Y g: X → Y h: Y → X r: f ~ g s: invmap f ~ h
remakeweqboth r s = g
X, Y: UU f: X ≃ Y g: X → Y h: Y → X r: f ~ g s: invmap f ~ h
remakeweqboth r s = g
X, Y: UU f: X ≃ Y g: X → Y h: Y → X r: f ~ g s: invmap f ~ h
remakeweqboth r s = g
apply idpath.Defined.
X, Y: UU f: X ≃ Y g: X → Y h: Y → X r: f ~ g s: invmap f ~ h
invmap (remakeweqboth r s) = h
X, Y: UU f: X ≃ Y g: X → Y h: Y → X r: f ~ g s: invmap f ~ h
invmap (remakeweqboth r s) = h
X, Y: UU f: X ≃ Y g: X → Y h: Y → X r: f ~ g s: invmap f ~ h
invmap (remakeweqboth r s) = h
apply idpath.Defined.
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f1 <-> isweq f2
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f1 <-> isweq f2
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f1 <-> isweq f2
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f1 → isweq f2
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f2 → isweq f1
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f1 → isweq f2
apply isweqhomot; assumption.
X, Y: UU f1, f2: X → Y h: f1 ~ f2
isweq f2 → isweq f1
apply isweqhomot, invhomot; assumption.Defined.
X: UU f, g: X → unit
isweq f → isweq g
X: UU f, g: X → unit
isweq f → isweq g
X: UU f, g: X → unit i: isweq f
isweq g
X: UU f, g: X → unit i: isweq f
f ~ g
X: UU f, g: X → unit i: isweq f h: f ~ g
isweq g
X: UU f, g: X → unit i: isweq f
f ~ g
X: UU f, g: X → unit i: isweq f t: X
f t = g t
apply isProofIrrelevantUnit.
X: UU f, g: X → unit i: isweq f h: f ~ g
isweq g
exact (isweqhomot f g h i).Defined.
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y
isweq f
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y
isweq f
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y
isweq f
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y
∏ y : Y, iscontr (hfiber f y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y
iscontr (hfiber f y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y
iscontr (hfiber (f ∘ g) y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y X0: iscontr (hfiber (f ∘ g) y)
iscontr (hfiber f y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y
∏ y : Y, y = f (g y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y efg': ∏ y : Y, y = f (g y)
iscontr (hfiber (f ∘ g) y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y X0: iscontr (hfiber (f ∘ g) y)
iscontr (hfiber f y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y
∏ y : Y, y = f (g y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y, y0: Y
y0 = f (g y0)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y, y0: Y
f (g y0) = y0
apply (efg y0).
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y efg': ∏ y : Y, y = f (g y)
iscontr (hfiber (f ∘ g) y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y X0: iscontr (hfiber (f ∘ g) y)
iscontr (hfiber f y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y X0: iscontr (hfiber (f ∘ g) y)
iscontr (hfiber f y)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x efg: ∏ y : Y, f (g y) = y y: Y X0: iscontr (hfiber (f ∘ g) y)
iscontr (hfiber (f ∘ g) y)
apply X0.Defined.(** This is kept to preserve compatibility with publications that use the name "gradth" for the "grad theorem". *)#[deprecated(note="Use isweq_iso instead.")]
Notationgradth := isweq_iso (only parsing).Definitionweq_iso {XY : UU} (f : X -> Y) (g : Y -> X)
(egf: ∏ x : X, g (f x) = x)
(efg: ∏ y : Y, f (g y) = y) : X ≃ Y :=
make_weq _ (isweq_iso _ _ egf efg).DefinitionUniqueConstruction {XY:UU} (f:X->Y) :=
(∏ y, ∑ x, f x = y) × (∏ x x', f x = f x' -> x = x').
X, Y: UU f: X → Y
UniqueConstruction f → isweq f
X, Y: UU f: X → Y
UniqueConstruction f → isweq f
X, Y: UU f: X → Y bij: UniqueConstruction f
isweq f
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y
isweq f
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij)
isweq f
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij)
Y → X
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij)
∏ x : X, ?g (f x) = x
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij)
∏ y : Y, f (?g y) = y
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij)
Y → X
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij) y: Y
X
exact (pr1 (sur y)).
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij)
∏ x : X, (λy : Y, pr1 (sur y)) (f x) = x
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij) x: X
(λy : Y, pr1 (sur y)) (f x) = x
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij) x: X
pr1 (sur (f x)) = x
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: ∏ x x' : X, f x = f x' → x = x' x: X
pr1 (sur (f x)) = x
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: ∏ x x' : X, f x = f x' → x = x' x: X
f (pr1 (sur (f x))) = f x
exact (pr2 (sur (f x))).
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij)
∏ y : Y, f ((λy0 : Y, pr1 (sur y0)) y) = y
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij) y: Y
f ((λy : Y, pr1 (sur y)) y) = y
X, Y: UU f: X → Y bij: UniqueConstruction f sur: ∏ y : Y, ∑ x : X, f x = y inj: (λ_ : ∏ y : Y, ∑ x : X, f x = y, ∏ x x' : X, f x = f x' → x = x') (pr1 bij) y: Y
X, Y: UU w: X ≃ Y efg: ∏ y : Y, w (invmap w y) = y
isweq (invmap w)
X, Y: UU w: X ≃ Y efg: ∏ y : Y, w (invmap w y) = y
isweq (invmap w)
X, Y: UU w: X ≃ Y efg: ∏ y : Y, w (invmap w y) = y
∏ x : X, invmap w (w x) = x
X, Y: UU w: X ≃ Y efg: ∏ y : Y, w (invmap w y) = y egf: ∏ x : X, invmap w (w x) = x
isweq (invmap w)
X, Y: UU w: X ≃ Y efg: ∏ y : Y, w (invmap w y) = y egf: ∏ x : X, invmap w (w x) = x
isweq (invmap w)
apply (isweq_iso _ _ efg egf).Defined.Definitioninvweq {XY : UU} (w : X ≃ Y) : Y ≃ X :=
make_weq (invmap w) (isweqinvmap w).
X, Y: UU w: X ≃ Y x: X
invmap (invweq w) x = w x
X, Y: UU w: X ≃ Y x: X
invmap (invweq w) x = w x
X, Y: UU w: X ≃ Y x: X
invmap (invweq w) x = w x
apply idpath.Defined.
X, Y: UU w: X ≃ Y
invweq w = invmap w
(* useful for rewriting *)
X, Y: UU w: X ≃ Y
invweq w = invmap w
X, Y: UU w: X ≃ Y
invweq w = invmap w
apply idpath.Defined.
X, Y: UU w: X ≃ Y is: iscontr X
iscontr Y
X, Y: UU w: X ≃ Y is: iscontr X
iscontr Y
X, Y: UU w: X ≃ Y is: iscontr X
iscontr Y
apply (iscontrweqb (invweq w) is).Defined.(** Equality between pairs is equivalent to pairs of equalities between components. Theorem adapted from HoTT library http://github.com/HoTT/HoTT. *)DefinitionPathPair {A : UU} {B : A -> UU} (xy : ∑ x, B x) :=
∑ p : pr1 x = pr1 y, transportf _ p (pr2 x) = pr2 y.Notation"a ╝ b" := (PathPair a b) : type_scope.(* the two horizontal lines represent an equality in the base and the two vertical lines represent an equality in the fiber *)(* in agda input mode use \--= and select the 6-th one in the first set, or use \chimney *)
A: UU B: A → UU x, y: ∑ x : A, B x
x = y ≃ x ╝ y
A: UU B: A → UU x, y: ∑ x : A, B x
x = y ≃ x ╝ y
A: UU B: A → UU x, y: ∑ x : A, B x
x = y ≃ x ╝ y
A: UU B: A → UU x, y: ∑ x : A, B x
isweq (λr : x = y, base_paths x y r,, fiber_paths r)
A: UU B: A → UU x, y: ∑ x : A, B x
∏ x0 : x = y,
total2_paths_f
(pr1 (base_paths x y x0,, fiber_paths x0))
(pr2 (base_paths x y x0,, fiber_paths x0)) = x0
A: UU B: A → UU x, y: ∑ x : A, B x
∏
y0 : ∑ p : pr1 x = pr1 y,
transportf (λx : A, B x) p (pr2 x) = pr2 y,
base_paths x y (total2_paths_f (pr1 y0) (pr2 y0)),,
fiber_paths (total2_paths_f (pr1 y0) (pr2 y0)) = y0
A: UU B: A → UU x, y: ∑ x : A, B x
∏ x0 : x = y,
total2_paths_f
(pr1 (base_paths x y x0,, fiber_paths x0))
(pr2 (base_paths x y x0,, fiber_paths x0)) = x0
A: UU B: A → UU x, y: ∑ x : A, B x p: x = y
total2_paths_f
(pr1 (base_paths x y p,, fiber_paths p))
(pr2 (base_paths x y p,, fiber_paths p)) = p
apply total2_fiber_paths.
A: UU B: A → UU x, y: ∑ x : A, B x
∏
y0 : ∑ p : pr1 x = pr1 y,
transportf (λx : A, B x) p (pr2 x) = pr2 y,
base_paths x y (total2_paths_f (pr1 y0) (pr2 y0)),,
fiber_paths (total2_paths_f (pr1 y0) (pr2 y0)) = y0
A: UU B: A → UU x, y: ∑ x : A, B x p: pr1 x = pr1 y q: transportf (λx : A, B x) p (pr2 x) = pr2 y
A: UU B: A → UU x, y: ∑ x : A, B x p: pr1 x = pr1 y q: transportf (λx : A, B x) p (pr2 x) = pr2 y
base_paths x y (total2_paths_f p q),,
fiber_paths (total2_paths_f p q) = p,, q
A: UU B: A → UU x, y: ∑ x : A, B x p: pr1 x = pr1 y q: transportf (λx : A, B x) p (pr2 x) = pr2 y
transportf
(λp : pr1 x = pr1 y,
transportf (λx : A, B x) p (pr2 x) = pr2 y)
(base_total2_paths q)
(fiber_paths (total2_paths_f p q)) = q
apply transportf_fiber_total2_paths.Defined.(** The standard weak equivalence from [ unit ] to a contractible type *)
X: UU is: iscontr X
unit ≃ X
X: UU is: iscontr X
unit ≃ X
X: UU is: iscontr X
unit ≃ X
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X
unit ≃ X
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit
unit ≃ X
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit
isweq f
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit
∏ t : unit, g (f t) = t
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit egf: ∏ t : unit, g (f t) = t
isweq f
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit
∏ t : unit, g (f t) = t
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit t: unit
g (f t) = t
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit
g (f tt) = tt
apply idpath.
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit egf: ∏ t : unit, g (f t) = t
isweq f
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit egf: ∏ t : unit, g (f t) = t
∏ x : X, f (g x) = x
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit egf: ∏ t : unit, g (f t) = t efg: ∏ x : X, f (g x) = x
isweq f
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit egf: ∏ t : unit, g (f t) = t
∏ x : X, f (g x) = x
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit egf: ∏ t : unit, g (f t) = t x: X
f (g x) = x
apply (! (pr2 is x)).
X: UU is: iscontr X f:= λ_ : unit, pr1 is: unit → X g:= λ_ : X, tt: X → unit egf: ∏ t : unit, g (f t) = t efg: ∏ x : X, f (g x) = x
isweq f
apply (isweq_iso _ _ egf efg).Defined.(** A weak equivalence between types defines weak equivalences on the corresponding [ paths ] types. *)
X, Y: UU w: X ≃ Y x, x': X
isweq (maponpaths w)
X, Y: UU w: X ≃ Y x, x': X
isweq (maponpaths w)
X, Y: UU w: X ≃ Y x, x': X
isweq (maponpaths w)
apply (isweq_iso (@maponpaths _ _ w x x')
(@invmaponpathsweq _ _ w x x')
(@pathsweq3 _ _ w x x')
(@pathsweq4 _ _ w x x')).Defined.Definitionweqonpaths {XY : UU} (w : X ≃ Y) (xx' : X) : x = x' ≃ w x = w x'
:= make_weq _ (isweqmaponpaths w x x').(** The inverse path and the composition with a path functions are weak equivalences *)
X: Type x, y: X
isweq pathsinv0
X: Type x, y: X
isweq pathsinv0
X: Type x, y: X
isweq pathsinv0
X: Type x, y: X p: y = x
iscontr (hfiber pathsinv0 p)
X: Type x, y: X p: y = x
hfiber pathsinv0 p
X: Type x, y: X p: y = x
(λcntr : hfiber pathsinv0 p,
∏ t : hfiber pathsinv0 p, t = cntr) ?pr1
X: Type x, y: X p: y = x
hfiber pathsinv0 p
X: Type x, y: X p: y = x
! (! p) = p
apply pathsinv0inv0.
X: Type x, y: X p: y = x
(λcntr : hfiber pathsinv0 p,
∏ t : hfiber pathsinv0 p, t = cntr)
(! p,, pathsinv0inv0 p)
X: Type x, y: X p: y = x
∏ t : hfiber pathsinv0 p, t = ! p,, pathsinv0inv0 p
reflexivity.Defined.Definitionweqpathsinv0 {X : UU} (xx' : X) : x = x' ≃ x' = x
:= make_weq _ (isweqpathsinv0 x x').
X: UU x, x', x'': X e': x' = x''
isweq (λe : x = x', e @ e')
X: UU x, x', x'': X e': x' = x''
isweq (λe : x = x', e @ e')
X: UU x, x', x'': X e': x' = x''
isweq (λe : x = x', e @ e')
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x''
isweq f
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x'
isweq f
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x'
∏ e : x = x', g (f e) = e
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x' egf: ∏ e : x = x', g (f e) = e
isweq f
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x'
∏ e : x = x', g (f e) = e
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x' e: x = x'
g (f e) = e
X: UU x, x'': X e': x = x'' f:= λe : x = x, e @ e': x = x → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x
g (f (idpath x)) = idpath x
X: UU x: X f:= λe : x = x, e @ idpath x: x = x → x = x g:= λe'' : x = x, e'' @ ! idpath x: x = x → x = x
g (f (idpath x)) = idpath x
apply idpath.
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x' egf: ∏ e : x = x', g (f e) = e
isweq f
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x' egf: ∏ e : x = x', g (f e) = e
∏ e : x = x'', f (g e) = e
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x' egf: ∏ e : x = x', g (f e) = e efg: ∏ e : x = x'', f (g e) = e
isweq f
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x' egf: ∏ e : x = x', g (f e) = e
∏ e : x = x'', f (g e) = e
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x' egf: ∏ e : x = x', g (f e) = e e: x = x''
f (g e) = e
X: UU x, x': X e': x' = x f:= λe : x = x', e @ e': x = x' → x = x g:= λe'' : x = x, e'' @ ! e': x = x → x = x' egf: ∏ e : x = x', g (f e) = e
X: UU x, x', x'': X e': x' = x'' f:= λe : x = x', e @ e': x = x' → x = x'' g:= λe'' : x = x'', e'' @ ! e': x = x'' → x = x' egf: ∏ e : x = x', g (f e) = e efg: ∏ e : x = x'', f (g e) = e
isweq f
apply (isweq_iso f g egf efg).Defined.(** Weak equivalences to and from coconuses and total path spaces *)
X, Y: UU f: X → Y
isweq (tococonusf f)
X, Y: UU f: X → Y
isweq (tococonusf f)
X, Y: UU f: X → Y
isweq (tococonusf f)
apply (isweq_iso _ _ (homotfromtococonusf f) (homottofromcoconusf f)).Defined.Definitionweqtococonusf {XY : UU} (f : X -> Y) : X ≃ coconusf f :=
make_weq _ (isweqtococonusf f).
X, Y: UU f: X → Y
isweq (fromcoconusf f)
X, Y: UU f: X → Y
isweq (fromcoconusf f)
X, Y: UU f: X → Y
isweq (fromcoconusf f)
apply (isweq_iso _ _ (homottofromcoconusf f) (homotfromtococonusf f)).Defined.Definitionweqfromcoconusf {XY : UU} (f : X -> Y) : coconusf f ≃ X :=
make_weq _ (isweqfromcoconusf f).
T: UU
isweq (deltap T)
T: UU
isweq (deltap T)
T: UU
isweq (deltap T)
T: UU ff:= deltap T: T → pathsspace T
isweq ff
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T
isweq ff
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T
∏ t : T, gg (ff t) = t
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t
isweq ff
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T
∏ t : T, gg (ff t) = t
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T t: T
gg (ff t) = t
apply idpath.
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t
isweq ff
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t
∏ tte : pathsspace T, ff (gg tte) = tte
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t efg: ∏ tte : pathsspace T, ff (gg tte) = tte
isweq ff
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t
∏ tte : pathsspace T, ff (gg tte) = tte
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t tte: pathsspace T
ff (gg tte) = tte
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t t: T c: coconusfromt T t
ff (gg (t,, c)) = t,, c
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t t, x: T e: t = x
ff (gg (t,, x,, e)) = t,, x,, e
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t t: T
ff (gg (t,, t,, idpath t)) = t,, t,, idpath t
apply idpath.
T: UU ff:= deltap T: T → pathsspace T gg:= λz : pathsspace T, pr1 z: pathsspace T → T egf: ∏ t : T, gg (ff t) = t efg: ∏ tte : pathsspace T, ff (gg tte) = tte
isweq ff
apply (isweq_iso _ _ egf efg).Defined.
T: UU
isweq (λa : pathsspace' T, pr1 (pr1 a))
T: UU
isweq (λa : pathsspace' T, pr1 (pr1 a))
T: UU
isweq (λa : pathsspace' T, pr1 (pr1 a))
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T
isweq f
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T
isweq f
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T
∏ t : T, f (g t) = t
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t
isweq f
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T
∏ t : T, f (g t) = t
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T t: T
f (g t) = t
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T t: T
pr1 (pr1 (g t)) = t
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T t: T
pr1 (pr1 (make_dirprod t t,, idpath t)) = t
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T t: T
t = t
apply idpath.
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t
isweq f
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t
∏ a : pathsspace' T, g (f a) = a
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t egf: ∏ a : pathsspace' T, g (f a) = a
isweq f
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t
∏ a : pathsspace' T, g (f a) = a
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t a: pathsspace' T
g (f a) = a
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t xy: T × T e: pr1 xy = pr2 xy
g (f (xy,, e)) = xy,, e
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t x, y: T e: pr1 (x,, y) = pr2 (x,, y)
g (f ((x,, y),, e)) = (x,, y),, e
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t x, y: T e: x = y
g (f ((x,, y),, e)) = (x,, y),, e
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t x: T
g (f ((x,, x),, idpath x)) = (x,, x),, idpath x
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t x: T
g (pr1 (pr1 ((x,, x),, idpath x))) =
(x,, x),, idpath x
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t x: T
T: UU f:= λa : pathsspace' T, pr1 (pr1 a): pathsspace' T → T g:= λt : T,
make_dirprod t t,, idpath t : pathsspace' T: T → pathsspace' T efg: ∏ t : T, f (g t) = t egf: ∏ a : pathsspace' T, g (f a) = a
isweq f
apply (isweq_iso _ _ egf efg).Defined.(** The weak equivalence between hfibers of homotopic functions *)
X, Y: UU f, g: X → Y h: f ~ g y: Y
hfiber f y ≃ hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y
hfiber f y ≃ hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y
hfiber f y ≃ hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y
hfiber f y ≃ hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y
hfiber f y ≃ hfiber g y
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y
isweq ff
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y
∏ xe : hfiber g y, ff (gg xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe
isweq ff
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y
∏ xe : hfiber g y, ff (gg xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y xe: hfiber g y
ff (gg xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y
ff (gg (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y
ff (gg (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y
! h x @ h x @ e = maponpaths g (idpath x) @ e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y eee: ! h x @ h x @ e = maponpaths g (idpath x) @ e
ff (gg (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y
! h x @ h x @ e = maponpaths g (idpath x) @ e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y
! h x @ h x @ e = e
X, Y: UU f, g: X → Y h: f ~ g x: X ff:= hfibershomotftog f g h (g x): hfiber f (g x) → hfiber g (g x) gg:= hfibershomotgtof f g h (g x): hfiber g (g x) → hfiber f (g x)
! h x @ h x @ idpath (g x) = idpath (g x)
X, Y: UU f, g: X → Y h: f ~ g x: X ff:= hfibershomotftog f g h (f x): hfiber f (f x) → hfiber g (f x) gg:= hfibershomotgtof f g h (f x): hfiber g (f x) → hfiber f (f x)
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y eee: ! h x @ h x @ e = maponpaths g (idpath x) @ e
ff (gg (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y eee: ! h x @ h x @ e = maponpaths g (idpath x) @ e xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y
ff (gg (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y x: X e: g x = y eee: ! h x @ h x @ e = maponpaths g (idpath x) @ e xe1:= make_hfiber g x (! h x @ h x @ e): hfiber g y xe2:= make_hfiber g x e: hfiber g y
ff (gg (x,, e)) = x,, e
apply (hfibertriangle2 g xe1 xe2 (idpath x) eee).
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe
isweq ff
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe
∏ xe : hfiber f y, gg (ff xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe eggff: ∏ xe : hfiber f y, gg (ff xe) = xe
isweq ff
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe
∏ xe : hfiber f y, gg (ff xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe xe: hfiber f y
gg (ff xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y
gg (ff (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y
gg (ff (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y
h x @ ! h x @ e = maponpaths f (idpath x) @ e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y eee: h x @ ! h x @ e = maponpaths f (idpath x) @ e
gg (ff (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y
h x @ ! h x @ e = maponpaths f (idpath x) @ e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y
h x @ ! h x @ e = e
X, Y: UU f, g: X → Y h: f ~ g x: X ff:= hfibershomotftog f g h (f x): hfiber f (f x) → hfiber g (f x) gg:= hfibershomotgtof f g h (f x): hfiber g (f x) → hfiber f (f x) effgg: ∏ xe : hfiber g (f x), ff (gg xe) = xe
h x @ ! h x @ idpath (f x) = idpath (f x)
X, Y: UU f, g: X → Y h: f ~ g x: X ff:= hfibershomotftog f g h (f x): hfiber f (f x) → hfiber g (f x) gg:= hfibershomotgtof f g h (f x): hfiber g (f x) → hfiber f (f x) effgg: ∏ xe : hfiber g (f x), ff (gg xe) = xe
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y eee: h x @ ! h x @ e = maponpaths f (idpath x) @ e
gg (ff (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y eee: h x @ ! h x @ e = maponpaths f (idpath x) @ e xe1:= make_hfiber f x (h x @ ! h x @ e): hfiber f y
gg (ff (x,, e)) = x,, e
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe x: X e: f x = y eee: h x @ ! h x @ e = maponpaths f (idpath x) @ e xe1:= make_hfiber f x (h x @ ! h x @ e): hfiber f y xe2:= make_hfiber f x e: hfiber f y
gg (ff (x,, e)) = x,, e
apply (hfibertriangle2 f xe1 xe2 (idpath x) eee).
X, Y: UU f, g: X → Y h: f ~ g y: Y ff:= hfibershomotftog f g h y: hfiber f y → hfiber g y gg:= hfibershomotgtof f g h y: hfiber g y → hfiber f y effgg: ∏ xe : hfiber g y, ff (gg xe) = xe eggff: ∏ xe : hfiber f y, gg (ff xe) = xe
isweq ff
apply (isweq_iso _ _ eggff effgg).Defined.(** *** The 2-out-of-3 property of weak equivalences Theorems showing that if any two of three functions f, g, gf are weak equivalences then so is the third - the 2-out-of-3 property. *)
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x)
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x)
∏ y : Y, f (invf y) = y
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) efinvf: ∏ y : Y, f (invf y) = y
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x)
∏ y : Y, f (invf y) = y
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) y: Y
f (invf y) = y
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) y: Y
g (f (invf y)) = g y
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) y: Y int1: g (f (invf y)) = g y
f (invf y) = y
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) y: Y
g (f (invf y)) = g y
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) y: Y
g (f ((invgf ∘ g) y)) = g y
apply (homotweqinvweq gfw (g y)).
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) y: Y int1: g (f (invf y)) = g y
f (invf y) = y
apply (invmaponpathsweq gw _ _ int1).
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) efinvf: ∏ y : Y, f (invf y) = y
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) efinvf: ∏ y : Y, f (invf y) = y
∏ x : X, invf (f x) = x
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) efinvf: ∏ y : Y, f (invf y) = y einvff: ∏ x : X, invf (f x) = x
isweq f
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) efinvf: ∏ y : Y, f (invf y) = y
∏ x : X, invf (f x) = x
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) efinvf: ∏ y : Y, f (invf y) = y x: X
invf (f x) = x
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) efinvf: ∏ y : Y, f (invf y) = y x: X
(invgf ∘ g) (f x) = x
apply (homotinvweqweq gfw x).
X, Y, Z: UU f: X → Y g: Y → Z isgf: isweq (g ∘ f) isg: isweq g gw:= make_weq g isg: Y ≃ Z gfw:= make_weq (g ∘ f) isgf: X ≃ Z invg:= invmap gw: Z → Y invgf:= invmap gfw: Z → X invf:= invgf ∘ g: ∏ x : Y, (λ_ : Z, X) (g x) efinvf: ∏ y : Y, f (invf y) = y einvff: ∏ x : X, invf (f x) = x
isweq f
apply (isweq_iso f invf einvff efinvf).Defined.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f)
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f)
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f)
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x)
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x)
∏ z : Z, g (invg z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x)
∏ z : Z, g (invg z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) z: Z
g (invg z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) z: Z
g ((f ∘ invgf) z) = z
apply (homotweqinvweq wgf z).
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z
∏ y : Y, invg (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z einvgg: ∏ y : Y, invg (g y) = y
isweq g
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z
∏ y : Y, invg (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y
invg (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y
isweq invf
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y
isweq invf
apply isweqinvmap.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf
isweq invgf
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf
isweq invgf
apply isweqinvmap.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf
g y = (g ∘ f) (invf y)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y)
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y)
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y)
(g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y)
(g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y)
(g ∘ f) (invgf (g y)) = g y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int3: (g ∘ f) (invgf (g y)) = g y
(g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y)
(g ∘ f) (invgf (g y)) = g y
apply (homotweqinvweq wgf).
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int3: (g ∘ f) (invgf (g y)) = g y
(g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int3: (g ∘ f) (invgf (g y)) = g y
(g ∘ f) (invgf (g y)) = g y
apply int3.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
invgf (g y) = invf y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
invgf (g y) = invf y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y)
wgf (invgf (g y)) = wgf (invf y)
apply int2.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y
invf (f (invgf (g y))) = invgf (g y)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y int5: invf (f (invgf (g y))) = invgf (g y)
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y
invf (f (invgf (g y))) = invgf (g y)
apply (homotinvweqweq wf).
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y int5: invf (f (invgf (g y))) = invgf (g y)
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y int5: invf (f (invgf (g y))) = invgf (g y)
invf (f (invgf (g y))) = invf y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y int5: invf (f (invgf (g y))) = invgf (g y) int6: invf (f (invgf (g y))) = invf y
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y int5: invf (f (invgf (g y))) = invgf (g y)
invf (f (invgf (g y))) = invf y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invgf (g y)) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invgf (g y)) int5: invf (f (invgf (g y))) = invgf (g y)
invf (f (invgf (g y))) = invgf (g y)
apply int5.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y int5: invf (f (invgf (g y))) = invgf (g y) int6: invf (f (invgf (g y))) = invf y
(f ∘ invgf) (g y) = y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y int5: invf (f (invgf (g y))) = invgf (g y) int6: invf (f (invgf (g y))) = invf y
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z y: Y isinvf: isweq invf isinvgf: isweq invgf int1: g y = (g ∘ f) (invf y) int2: (g ∘ f) (invgf (g y)) = (g ∘ f) (invf y) int4: invgf (g y) = invf y int5: invf (f (invgf (g y))) = invgf (g y) int6: invf (f (invgf (g y))) = invf y
invf (f (invgf (g y))) = invf y
apply int6.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isgf: isweq (g ∘ f) wf:= make_weq f isf: X ≃ Y wgf:= make_weq (g ∘ f) isgf: X ≃ Z invf:= invmap wf: Y → X invgf:= invmap wgf: Z → X invg:= f ∘ invgf: ∏ x : Z, (λ_ : X, Y) (invgf x) eginvg: ∏ z : Z, g (invg z) = z einvgg: ∏ y : Y, invg (g y) = y
isweq g
apply (isweq_iso g invg einvgg eginvg).Defined.
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x
isweq f → isweq g
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x
isweq f → isweq g
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x w: isweq f
isweq g
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x w: isweq f
isweq (g ∘ f)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x w: isweq f int1: isweq (g ∘ f)
isweq g
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x w: isweq f
isweq (g ∘ f)
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x w: isweq f
isweq (idfun X)
apply idisweq.
X, Y: UU f: X → Y g: Y → X egf: ∏ x : X, g (f x) = x w: isweq f int1: isweq (g ∘ f)
isweq g
apply (twooutof3b f g w int1).Defined.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g
isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g
isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g
isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y
isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z
isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X
isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y
isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x)
isweq gf
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x)
isweq gf
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x)
∏ x : X, invgf (gf x) = x
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x
isweq gf
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x)
∏ x : X, invgf (gf x) = x
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X
invgf (gf x) = x
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X
invf (invg (g (f x))) = invf (f x)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X int1: invf (invg (g (f x))) = invf (f x)
invgf (gf x) = x
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X
invf (invg (g (f x))) = invf (f x)
apply (maponpaths invf (homotinvweqweq wg (f x))).
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X int1: invf (invg (g (f x))) = invf (f x)
invgf (gf x) = x
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X int1: invf (invg (g (f x))) = invf (f x)
invf (f x) = x
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X int1: invf (invg (g (f x))) = invf (f x) int2: invf (f x) = x
invgf (gf x) = x
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X int1: invf (invg (g (f x))) = invf (f x)
invf (f x) = x
apply (homotinvweqweq wf x).
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X int1: invf (invg (g (f x))) = invf (f x) int2: invf (f x) = x
invgf (gf x) = x
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) x: X int2: invf (invg (g (f x))) = x
invgf (gf x) = x
apply int2.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x
isweq gf
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x
∏ z : Z, gf (invgf z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x einvgfgf: ∏ z : Z, gf (invgf z) = z
isweq gf
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x
∏ z : Z, gf (invgf z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z
gf (invgf z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z
g (f (invgf z)) = g (invg z)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z int1: g (f (invgf z)) = g (invg z)
gf (invgf z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z
g (f (invgf z)) = g (invg z)
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z
g (f ((invf ∘ invg) z)) = g (invg z)
apply (maponpaths g (homotweqinvweq wf (invg z))).
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z int1: g (f (invgf z)) = g (invg z)
gf (invgf z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z int1: g (f (invgf z)) = g (invg z)
g (invg z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z int1: g (f (invgf z)) = g (invg z) int2: g (invg z) = z
gf (invgf z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z int1: g (f (invgf z)) = g (invg z)
g (invg z) = z
apply (homotweqinvweq wg z).
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z int1: g (f (invgf z)) = g (invg z) int2: g (invg z) = z
gf (invgf z) = z
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x z: Z int2: g (f (invgf z)) = z
gf (invgf z) = z
apply int2.
X, Y, Z: UU f: X → Y g: Y → Z isf: isweq f isg: isweq g wf:= make_weq f isf: X ≃ Y wg:= make_weq g isg: Y ≃ Z invf:= invmap wf: Y → X invg:= invmap wg: Z → Y gf:= g ∘ f: ∏ x : X, (λ_ : Y, Z) (f x) invgf:= invf ∘ invg: ∏ x : Z, (λ_ : Y, X) (invg x) egfinvgf: ∏ x : X, invgf (gf x) = x einvgfgf: ∏ z : Z, gf (invgf z) = z
X, Y, Z: UU f: X → Y g: Y → Z i: isweq f j: isweq g
isweq (g ∘ f)
exact (twooutof3c f g i j).
X, Y, Z: UU f: X → Y g: Y → Z i: isweq f
isweq (g ∘ f) → isweq g
X, Y, Z: UU f: X → Y g: Y → Z i: isweq f j: isweq (g ∘ f)
isweq g
exact (twooutof3b f g i j).Defined.
X, Y, Z: UU f: X → Y g: Y → Z
isweq g → isweq f <-> isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z
isweq g → isweq f <-> isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z i: isweq g
isweq f <-> isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z i: isweq g
isweq f → isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z i: isweq g
isweq (g ∘ f) → isweq f
X, Y, Z: UU f: X → Y g: Y → Z i: isweq g
isweq f → isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z i: isweq g j: isweq f
isweq (g ∘ f)
exact (twooutof3c f g j i).
X, Y, Z: UU f: X → Y g: Y → Z i: isweq g
isweq (g ∘ f) → isweq f
X, Y, Z: UU f: X → Y g: Y → Z i: isweq g j: isweq (g ∘ f)
isweq f
exact (twooutof3a f g j i).Defined.
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z
g ∘ f ~ h → isweq g → isweq f <-> isweq h
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z
g ∘ f ~ h → isweq g → isweq f <-> isweq h
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq g
isweq f <-> isweq h
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq g
isweq f <-> isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq g
isweq (g ∘ f) <-> isweq h
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq g
isweq f <-> isweq (g ∘ f)
apply twooutof3c_iff_1; assumption.
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq g
isweq (g ∘ f) <-> isweq h
apply isweqhomot_iff; assumption.Defined.
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z
g ∘ f ~ h → isweq f → isweq g <-> isweq h
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z
g ∘ f ~ h → isweq f → isweq g <-> isweq h
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq f
isweq g <-> isweq h
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq f
isweq g <-> isweq (g ∘ f)
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq f
isweq (g ∘ f) <-> isweq h
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq f
isweq g <-> isweq (g ∘ f)
apply twooutof3c_iff_2; assumption.
X, Y, Z: UU f: X → Y g: Y → Z h: X → Z r: g ∘ f ~ h i: isweq f
isweq (g ∘ f) <-> isweq h
apply isweqhomot_iff; assumption.Defined.(** *** Any function between contractible types is a weak equivalence *)
X, Y: UU f: X → Y isx: iscontr X isy: iscontr Y
isweq f
X, Y: UU f: X → Y isx: iscontr X isy: iscontr Y
isweq f
X, Y: UU f: X → Y isx: iscontr X isy: iscontr Y
isweq f
X, Y: UU f: X → Y isx: iscontr X isy: iscontr Y py:= λ_ : Y, tt: Y → unit
isweq f
apply (twooutof3a f py (isweqcontrtounit isx) (isweqcontrtounit isy)).Defined.Definitionweqcontrcontr {XY : UU} (isx : iscontr X) (isy : iscontr Y) : X ≃ Y
:= make_weq (λ_, pr1 isy) (isweqcontrcontr _ isx isy).(** *** Composition of weak equivalences *)Definitionweqcomp {XYZ : UU} (w1 : X ≃ Y) (w2 : Y ≃ Z) : X ≃ Z :=
make_weq (λ (x : X), w2 (w1 x)) (twooutof3c w1 w2 (pr2 w1) (pr2 w2)).Declare Scope weq_scope.Notation"g ∘ f" := (weqcomp f g) : weq_scope.Delimit Scope weq_scope with weq.Ltacintermediate_weq Y' := apply (weqcomp (Y := Y')).
X, Y, Z: UU x: X f: X ≃ Y g: Y ≃ Z
(g ∘ f)%weq x = g (f x)
X, Y, Z: UU x: X f: X ≃ Y g: Y ≃ Z
(g ∘ f)%weq x = g (f x)
X, Y, Z: UU x: X f: X ≃ Y g: Y ≃ Z
(g ∘ f)%weq x = g (f x)
apply idpath.Defined.
X, Y, Z: UU f: X ≃ Y g: Y ≃ Z
(g ∘ f)%weq = g ∘ f
X, Y, Z: UU f: X ≃ Y g: Y ≃ Z
(g ∘ f)%weq = g ∘ f
X, Y, Z: UU f: X ≃ Y g: Y ≃ Z
(g ∘ f)%weq = g ∘ f
apply idpath.Defined.
X, Y, Z: UU f: X ≃ Y g: Y ≃ Z
invmap (g ∘ f)%weq = invmap f ∘ invmap g
X, Y, Z: UU f: X ≃ Y g: Y ≃ Z
invmap (g ∘ f)%weq = invmap f ∘ invmap g
X, Y, Z: UU f: X ≃ Y g: Y ≃ Z
invmap (g ∘ f)%weq = invmap f ∘ invmap g
apply idpath.Defined.(** *** The 2-out-of-6 (two-out-of-six) property of weak equivalences *)
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x)
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x) hupu:= homotinvweqweq (make_weq (v ∘ u) isuv)
:
pu ∘ u ~ idfun X: pu ∘ u ~ idfun X
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x) hupu:= homotinvweqweq (make_weq (v ∘ u) isuv)
:
pu ∘ u ~ idfun X: pu ∘ u ~ idfun X invvw:= invmap (make_weq (w ∘ v) isvw): K → Y
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x) hupu:= homotinvweqweq (make_weq (v ∘ u) isuv)
:
pu ∘ u ~ idfun X: pu ∘ u ~ idfun X invvw:= invmap (make_weq (w ∘ v) isvw): K → Y pv:= invvw ∘ w: ∏ x : Z, (λ_ : K, Y) (w x)
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x) hupu:= homotinvweqweq (make_weq (v ∘ u) isuv)
:
pu ∘ u ~ idfun X: pu ∘ u ~ idfun X invvw:= invmap (make_weq (w ∘ v) isvw): K → Y pv:= invvw ∘ w: ∏ x : Z, (λ_ : K, Y) (w x) hvpv:= homotinvweqweq (make_weq (w ∘ v) isvw)
:
pv ∘ v ~ idfun Y: pv ∘ v ~ idfun Y
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x) hupu:= homotinvweqweq (make_weq (v ∘ u) isuv)
:
pu ∘ u ~ idfun X: pu ∘ u ~ idfun X invvw:= invmap (make_weq (w ∘ v) isvw): K → Y pv:= invvw ∘ w: ∏ x : Z, (λ_ : K, Y) (w x) hvpv:= homotinvweqweq (make_weq (w ∘ v) isvw)
:
pv ∘ v ~ idfun Y: pv ∘ v ~ idfun Y h0:= funhomot v
(homotweqinvweq (make_weq (v ∘ u) isuv)): (λy : Z,
make_weq (v ∘ u) isuv
(invmap (make_weq (v ∘ u) isuv) y)) ∘ v ~
(λy : Z, y) ∘ v
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x) hupu:= homotinvweqweq (make_weq (v ∘ u) isuv)
:
pu ∘ u ~ idfun X: pu ∘ u ~ idfun X invvw:= invmap (make_weq (w ∘ v) isvw): K → Y pv:= invvw ∘ w: ∏ x : Z, (λ_ : K, Y) (w x) hvpv:= homotinvweqweq (make_weq (w ∘ v) isvw)
:
pv ∘ v ~ idfun Y: pv ∘ v ~ idfun Y h0:= funhomot v
(homotweqinvweq (make_weq (v ∘ u) isuv)): (λy : Z,
make_weq (v ∘ u) isuv
(invmap (make_weq (v ∘ u) isuv) y)) ∘ v ~
(λy : Z, y) ∘ v h1:= funhomot (u ∘ pu) (invhomot hvpv): idfun Y ∘ (u ∘ pu) ~ pv ∘ v ∘ (u ∘ pu)
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x) hupu:= homotinvweqweq (make_weq (v ∘ u) isuv)
:
pu ∘ u ~ idfun X: pu ∘ u ~ idfun X invvw:= invmap (make_weq (w ∘ v) isvw): K → Y pv:= invvw ∘ w: ∏ x : Z, (λ_ : K, Y) (w x) hvpv:= homotinvweqweq (make_weq (w ∘ v) isvw)
:
pv ∘ v ~ idfun Y: pv ∘ v ~ idfun Y h0:= funhomot v
(homotweqinvweq (make_weq (v ∘ u) isuv)): (λy : Z,
make_weq (v ∘ u) isuv
(invmap (make_weq (v ∘ u) isuv) y)) ∘ v ~
(λy : Z, y) ∘ v h1:= funhomot (u ∘ pu) (invhomot hvpv): idfun Y ∘ (u ∘ pu) ~ pv ∘ v ∘ (u ∘ pu) h2:= homotfun h0 pv: pv
∘ ((λy : Z,
make_weq (v ∘ u) isuv
(invmap (make_weq (v ∘ u) isuv) y)) ∘ v) ~
pv ∘ ((λy : Z, y) ∘ v)
isweq u
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v) invuv:= invmap (make_weq (v ∘ u) isuv): Z → X pu:= invuv ∘ v: ∏ x : Y, (λ_ : Z, X) (v x) hupu:= homotinvweqweq (make_weq (v ∘ u) isuv)
:
pu ∘ u ~ idfun X: pu ∘ u ~ idfun X invvw:= invmap (make_weq (w ∘ v) isvw): K → Y pv:= invvw ∘ w: ∏ x : Z, (λ_ : K, Y) (w x) hvpv:= homotinvweqweq (make_weq (w ∘ v) isvw)
:
pv ∘ v ~ idfun Y: pv ∘ v ~ idfun Y h0:= funhomot v
(homotweqinvweq (make_weq (v ∘ u) isuv)): (λy : Z,
make_weq (v ∘ u) isuv
(invmap (make_weq (v ∘ u) isuv) y)) ∘ v ~
(λy : Z, y) ∘ v h1:= funhomot (u ∘ pu) (invhomot hvpv): idfun Y ∘ (u ∘ pu) ~ pv ∘ v ∘ (u ∘ pu) h2:= homotfun h0 pv: pv
∘ ((λy : Z,
make_weq (v ∘ u) isuv
(invmap (make_weq (v ∘ u) isuv) y)) ∘ v) ~
pv ∘ ((λy : Z, y) ∘ v) hpuu:= homotcomp (homotcomp h1 h2) hvpv: idfun Y ∘ (u ∘ pu) ~ idfun Y
isweq u
exact (isweq_iso u pu hupu hpuu).Defined.
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
isweq v
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
isweq v
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
isweq w
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
isweq w
X, Y, Z, K: UU u: X → Y v: Y → Z w: Z → K isuv: isweq (v ∘ u) isvw: isweq (w ∘ v)
isweq w
exact (twooutof3b _ _ (twooutofsixv isuv isvw) isvw).Defined.(** *** Pairwise direct products of weak equivalences *)
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y'
isweq (dirprodf w w')
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y'
isweq (dirprodf w w')
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y'
isweq (dirprodf w w')
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y'
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X'
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X'
∏ a : X × X', g (f a) = a
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' a: X × X'
g (f a) = a
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' x: X x': X'
g (f (x,, x')) = x,, x'
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' x: X x': X'
g (f (x,, x')) = x,, x'
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' x: X x': X'
invweq w (pr1 (f (x,, x'))) = x
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' x: X x': X'
invweq w' (pr2 (f (x,, x'))) = x'
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' x: X x': X'
invweq w' (pr2 (f (x,, x'))) = x'
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a
∏ a : Y × Y', f (g a) = a
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a efg: ∏ a : Y × Y', f (g a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a a: Y × Y'
f (g a) = a
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a efg: ∏ a : Y × Y', f (g a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a x: Y x': Y'
f (g (x,, x')) = x,, x'
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a efg: ∏ a : Y × Y', f (g a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a x: Y x': Y'
f (g (x,, x')) = x,, x'
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a efg: ∏ a : Y × Y', f (g a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a x: Y x': Y'
w (pr1 (g (x,, x'))) = x
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a x: Y x': Y'
w' (pr2 (g (x,, x'))) = x'
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a efg: ∏ a : Y × Y', f (g a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a x: Y x': Y'
w' (pr2 (g (x,, x'))) = x'
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a efg: ∏ a : Y × Y', f (g a) = a
isweq f
X, Y, X', Y': UU w: X ≃ Y w': X' ≃ Y' f:= dirprodf w w': X × X' → Y × Y' g:= dirprodf (invweq w) (invweq w'): Y × Y' → X × X' egf: ∏ a : X × X', g (f a) = a efg: ∏ a : Y × Y', f (g a) = a
isweq f
apply (isweq_iso _ _ egf efg).Defined.Definitionweqdirprodf {XYX'Y' : UU} (w : X ≃ Y) (w' : X' ≃ Y') : X × X' ≃ Y × Y'
:= make_weq _ (isweqdirprodf w w').(** *** Weak equivalence of a type and its direct product with the unit *)
X: UU
X ≃ X × unit
X: UU
X ≃ X × unit
X: UU
X ≃ X × unit
X: UU f:= λx : X, make_dirprod x tt: X → X × unit
X ≃ X × unit
X: UU f:= λx : X, make_dirprod x tt: X → X × unit
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X
∏ x : X, g (f x) = x
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X x: X
g (f x) = x
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x
∏ xu : X × unit, f (g xu) = xu
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x efg: ∏ xu : X × unit, f (g xu) = xu
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x xu: X × unit
f (g xu) = xu
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x efg: ∏ xu : X × unit, f (g xu) = xu
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x t: X x: unit
f (g (t,, x)) = t,, x
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x efg: ∏ xu : X × unit, f (g xu) = xu
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x t: X
f (g (t,, tt)) = t,, tt
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x efg: ∏ xu : X × unit, f (g xu) = xu
isweq f
X: UU f:= λx : X, make_dirprod x tt: X → X × unit g:= λxu : X × unit, pr1 xu: X × unit → X egf: ∏ x : X, g (f x) = x efg: ∏ xu : X × unit, f (g xu) = xu
isweq f
apply (isweq_iso f g egf efg).Defined.(** *** Associativity of [ total2 ] as a weak equivalence *)
X: UU P: X → UU Q: (∑ y, P y) → UU
(∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p)
X: UU P: X → UU Q: (∑ y, P y) → UU
(∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p)
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ y, Q y
∑ (x : X) (p : P x), Q (x,, p)
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ y, Q y
∑ p : P (pr1 (pr1 xpq)), Q (pr1 (pr1 xpq),, p)
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ y, Q y
Q (pr1 (pr1 xpq),, pr2 (pr1 xpq))
exact (pr2 xpq).Defined.
X: UU P: X → UU Q: (∑ y, P y) → UU
(∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y
X: UU P: X → UU Q: (∑ y, P y) → UU
(∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
∑ y, Q y
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
∑ y, P y
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
Q ?pr1
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
∑ y, P y
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
X
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
P ?pr1
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
X
apply (pr1 xpq).
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
P (pr1 xpq)
apply (pr1 (pr2 xpq)).
X: UU P: X → UU Q: (∑ y, P y) → UU xpq: ∑ (x : X) (p : P x), Q (x,, p)
Q (pr1 xpq,, pr1 (pr2 xpq))
exact (pr2 (pr2 xpq)).Defined.
X: UU P: X → UU Q: (∑ y, P y) → UU
(∑ y, Q y) ≃ (∑ (x : X) (p : P x), Q (x,, p))
X: UU P: X → UU Q: (∑ y, P y) → UU
(∑ y, Q y) ≃ (∑ (x : X) (p : P x), Q (x,, p))
X: UU P: X → UU Q: (∑ y, P y) → UU
(∑ y, Q y) ≃ (∑ (x : X) (p : P x), Q (x,, p))
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p)
(∑ y, Q y) ≃ (∑ (x : X) (p : P x), Q (x,, p))
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y
(∑ y, Q y) ≃ (∑ (x : X) (p : P x), Q (x,, p))
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y
isweq f
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y
∏ xpq : ∑ y, Q y, g (f xpq) = xpq
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y egf: ∏ xpq : ∑ y, Q y, g (f xpq) = xpq
isweq f
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y xpq: ∑ y, Q y
g (f xpq) = xpq
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y egf: ∏ xpq : ∑ y, Q y, g (f xpq) = xpq
isweq f
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y egf: ∏ xpq : ∑ y, Q y, g (f xpq) = xpq
isweq f
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y egf: ∏ xpq : ∑ y, Q y, g (f xpq) = xpq
∏ xpq : ∑ (x : X) (p : P x), Q (x,, p),
f (g xpq) = xpq
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y egf: ∏ xpq : ∑ y, Q y, g (f xpq) = xpq efg: ∏ xpq : ∑ (x : X) (p : P x), Q (x,, p), f (g xpq) = xpq
isweq f
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y egf: ∏ xpq : ∑ y, Q y, g (f xpq) = xpq xpq: ∑ (x : X) (p : P x), Q (x,, p)
f (g xpq) = xpq
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y egf: ∏ xpq : ∑ y, Q y, g (f xpq) = xpq efg: ∏ xpq : ∑ (x : X) (p : P x), Q (x,, p), f (g xpq) = xpq
isweq f
X: UU P: X → UU Q: (∑ y, P y) → UU f:= total2asstor P Q: (∑ y, Q y) → ∑ (x : X) (p : P x), Q (x,, p) g:= total2asstol P Q: (∑ (x : X) (p : P x), Q (x,, p)) → ∑ y, Q y egf: ∏ xpq : ∑ y, Q y, g (f xpq) = xpq efg: ∏ xpq : ∑ (x : X) (p : P x), Q (x,, p), f (g xpq) = xpq
isweq f
apply (isweq_iso _ _ egf efg).Defined.Definitionweqtotal2asstol {X : UU} (P : X -> UU) (Q : total2 P -> UU) :
(∑ x : X, ∑ p : P x, Q (tpair P x p)) ≃ total2 Q
:= invweq (weqtotal2asstor P Q).(** *** Associativity and commutativity of direct products as weak equivalences *)
X, Y, Z: UU
(X × Y) × Z ≃ X × Y × Z
X, Y, Z: UU
(X × Y) × Z ≃ X × Y × Z
X, Y, Z: UU
(X × Y) × Z ≃ X × Y × Z
apply weqtotal2asstor.Defined.Definitionweqdirprodasstol (XYZ : UU) : X × (Y × Z) ≃ (X × Y) × Z
:= invweq (weqdirprodasstor X Y Z).
X, Y: UU
X × Y ≃ Y × X
X, Y: UU
X × Y ≃ Y × X
X, Y: UU
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y
∏ xy : X × Y, g (f xy) = xy
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y xy: X × Y
g (f xy) = xy
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y,
make_dirprod (Preamble.pr2 xy) (Preamble.pr1 xy): X × Y → Y × X g:= λyx : Y × X,
make_dirprod (Preamble.pr2 yx) (Preamble.pr1 yx): Y × X → X × Y pr1: X pr2: Y
g (f (pr1,, pr2)) = pr1,, pr2
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy
∏ yx : Y × X, f (g yx) = yx
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy efg: ∏ yx : Y × X, f (g yx) = yx
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy yx: Y × X
f (g yx) = yx
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy efg: ∏ yx : Y × X, f (g yx) = yx
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y,
make_dirprod (Preamble.pr2 xy) (Preamble.pr1 xy): X × Y → Y × X g:= λyx : Y × X,
make_dirprod (Preamble.pr2 yx) (Preamble.pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy pr1: Y pr2: X
f (g (pr1,, pr2)) = pr1,, pr2
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy efg: ∏ yx : Y × X, f (g yx) = yx
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy efg: ∏ yx : Y × X, f (g yx) = yx
X × Y ≃ Y × X
X, Y: UU f:= λxy : X × Y, make_dirprod (pr2 xy) (pr1 xy): X × Y → Y × X g:= λyx : Y × X, make_dirprod (pr2 yx) (pr1 yx): Y × X → X × Y egf: ∏ xy : X × Y, g (f xy) = xy efg: ∏ yx : Y × X, f (g yx) = yx
isweq f
apply (isweq_iso _ _ egf efg).Defined.
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy)
≃ (∑ xy : Y × X, P (weqdirprodcomm Y X xy))
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy)
≃ (∑ xy : Y × X, P (weqdirprodcomm Y X xy))
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy)
≃ (∑ xy : Y × X, P (weqdirprodcomm Y X xy))
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy)
→ ∑ xy : Y × X, P (weqdirprodcomm Y X xy)
X, Y: UU P: X × Y → UU
(∑ xy : Y × X, P (weqdirprodcomm Y X xy))
→ ∑ xy : X × Y, P xy
X, Y: UU P: X × Y → UU
∏ x : ∑ xy : X × Y, P xy, ?g (?f x) = x
X, Y: UU P: X × Y → UU
∏ y : ∑ xy : Y × X, P (weqdirprodcomm Y X xy),
?f (?g y) = y
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy)
→ ∑ xy : Y × X, P (weqdirprodcomm Y X xy)
X, Y: UU P: X × Y → UU xyp: ∑ xy : X × Y, P xy
∑ xy : Y × X, P (weqdirprodcomm Y X xy)
exact ((pr2 (pr1 xyp),, pr1 (pr1 xyp)),,pr2 xyp).
X, Y: UU P: X × Y → UU
(∑ xy : Y × X, P (weqdirprodcomm Y X xy))
→ ∑ xy : X × Y, P xy
X, Y: UU P: X × Y → UU yxp: ∑ xy : Y × X, P (weqdirprodcomm Y X xy)
∏ x : ∑ xy : X × Y, P xy,
(λyxp : ∑ xy : Y × X, P (weqdirprodcomm Y X xy),
(pr2 (pr1 yxp),, pr1 (pr1 yxp)),, pr2 yxp)
((λxyp : ∑ xy : X × Y, P xy,
(pr2 (pr1 xyp),, pr1 (pr1 xyp)),, pr2 xyp) x) = x
X, Y: UU P: X × Y → UU xyp: ∑ xy : X × Y, P xy
(λyxp : ∑ xy : Y × X, P (weqdirprodcomm Y X xy),
(pr2 (pr1 yxp),, pr1 (pr1 yxp)),, pr2 yxp)
((λxyp : ∑ xy : X × Y, P xy,
(pr2 (pr1 xyp),, pr1 (pr1 xyp)),, pr2 xyp) xyp) =
xyp
apply idpath.
X, Y: UU P: X × Y → UU
∏ y : ∑ xy : Y × X, P (weqdirprodcomm Y X xy),
(λxyp : ∑ xy : X × Y, P xy,
(pr2 (pr1 xyp),, pr1 (pr1 xyp)),, pr2 xyp)
((λyxp : ∑ xy : Y × X, P (weqdirprodcomm Y X xy),
(pr2 (pr1 yxp),, pr1 (pr1 yxp)),, pr2 yxp) y) = y
X, Y: UU P: X × Y → UU yxp: ∑ xy : Y × X, P (weqdirprodcomm Y X xy)
(λxyp : ∑ xy : X × Y, P xy,
(pr2 (pr1 xyp),, pr1 (pr1 xyp)),, pr2 xyp)
((λyxp : ∑ xy : Y × X, P (weqdirprodcomm Y X xy),
(pr2 (pr1 yxp),, pr1 (pr1 yxp)),, pr2 yxp) yxp) =
yxp
apply idpath.Defined.
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy) ≃ (∑ (x : X) (y : Y), P (x,, y))
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy) ≃ (∑ (x : X) (y : Y), P (x,, y))
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy) → ∑ (x : X) (y : Y), P (x,, y)
X, Y: UU P: X × Y → UU
(∑ (x : X) (y : Y), P (x,, y)) → ∑ xy : X × Y, P xy
X, Y: UU P: X × Y → UU
∏ x : ∑ xy : X × Y, P xy, ?g (?f x) = x
X, Y: UU P: X × Y → UU
∏ y : ∑ (x : X) (y : Y), P (x,, y), ?f (?g y) = y
X, Y: UU P: X × Y → UU
(∑ xy : X × Y, P xy) → ∑ (x : X) (y : Y), P (x,, y)
X, Y: UU P: X × Y → UU xyp: ∑ xy : X × Y, P xy
∑ (x : X) (y : Y), P (x,, y)
exact (pr1 (pr1 xyp),,pr2 (pr1 xyp),, pr2 xyp).
X, Y: UU P: X × Y → UU
(∑ (x : X) (y : Y), P (x,, y)) → ∑ xy : X × Y, P xy
X, Y: UU P: X × Y → UU xyp: ∑ (x : X) (y : Y), P (x,, y)
∏ x : ∑ w : ∑ x : X, P x, Q (pr1 w),
(λxqp : ∑ w : ∑ x0 : X, Q x0, P (pr1 w),
(pr1 (pr1 xqp),, pr2 xqp),, pr2 (pr1 xqp))
((λxpq : ∑ w : ∑ x0 : X, P x0, Q (pr1 w),
(pr1 (pr1 xpq),, pr2 xpq),, pr2 (pr1 xpq)) x) = x
X: Type P, Q: X → UU x: ∑ w : ∑ x : X, P x, Q (pr1 w)
(λxqp : ∑ w : ∑ x : X, Q x, P (pr1 w),
(pr1 (pr1 xqp),, pr2 xqp),, pr2 (pr1 xqp))
((λxpq : ∑ w : ∑ x : X, P x, Q (pr1 w),
(pr1 (pr1 xpq),, pr2 xpq),, pr2 (pr1 xpq)) x) = x
apply idpath.
X: Type P, Q: X → UU
∏ y : ∑ w : ∑ x : X, Q x, P (pr1 w),
(λxpq : ∑ w : ∑ x : X, P x, Q (pr1 w),
(pr1 (pr1 xpq),, pr2 xpq),, pr2 (pr1 xpq))
((λxqp : ∑ w : ∑ x : X, Q x, P (pr1 w),
(pr1 (pr1 xqp),, pr2 xqp),, pr2 (pr1 xqp)) y) = y
X: Type P, Q: X → UU y: ∑ w : ∑ x : X, Q x, P (pr1 w)
(λxpq : ∑ w : ∑ x : X, P x, Q (pr1 w),
(pr1 (pr1 xpq),, pr2 xpq),, pr2 (pr1 xpq))
((λxqp : ∑ w : ∑ x : X, Q x, P (pr1 w),
(pr1 (pr1 xqp),, pr2 xqp),, pr2 (pr1 xqp)) y) = y
apply idpath.Defined.(** ** Binary coproducts and their basic properties *)(** Binary coproducts have not been introduced or used earlier except for thelines in the Preamble.v that define [ coprod ] and [ ⨿ ] as a notation forthe inductive type family [ sum ] that is defined in Coq.Init. *)(** *** Distributivity of coproducts and direct products as a weak equivalence *)
X, Y, Z: UU
X × Y ⨿ Z → (X × Y) ⨿ (X × Z)
X, Y, Z: UU
X × Y ⨿ Z → (X × Y) ⨿ (X × Z)
X, Y, Z: UU X0: X × Y ⨿ Z
(X × Y) ⨿ (X × Z)
X, Y, Z: UU t: X x: Y ⨿ Z
(X × Y) ⨿ (X × Z)
X, Y, Z: UU t: X y: Y
(X × Y) ⨿ (X × Z)
X, Y, Z: UU t: X z: Z
(X × Y) ⨿ (X × Z)
X, Y, Z: UU t: X z: Z
(X × Y) ⨿ (X × Z)
apply (ii2 (make_dirprod t z)).Defined.
X, Y, Z: UU
(X × Y) ⨿ (X × Z) → X × Y ⨿ Z
X, Y, Z: UU
(X × Y) ⨿ (X × Z) → X × Y ⨿ Z
X, Y, Z: UU X0: (X × Y) ⨿ (X × Z)
X × Y ⨿ Z
X, Y, Z: UU d: X × Y
X × Y ⨿ Z
X, Y, Z: UU d: X × Z
X × Y ⨿ Z
X, Y, Z: UU t: X x: Y
X × Y ⨿ Z
X, Y, Z: UU d: X × Z
X × Y ⨿ Z
X, Y, Z: UU d: X × Z
X × Y ⨿ Z
X, Y, Z: UU t: X x: Z
X × Y ⨿ Z
apply (make_dirprod t (ii2 x)).Defined.
X, Y, Z: UU
isweq (rdistrtoprod X Y Z)
X, Y, Z: UU
isweq (rdistrtoprod X Y Z)
X, Y, Z: UU
isweq (rdistrtoprod X Y Z)
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z)
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z)
∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) a: (X × Y) ⨿ (X × Z)
g (f a) = a
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) d: X × Y
g (f (inl d)) = inl d
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) d: X × Z
g (f (inr d)) = inr d
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) pr1: X pr2: Y
g (f (inl (pr1,, pr2))) = inl (pr1,, pr2)
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) d: X × Z
g (f (inr d)) = inr d
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) d: X × Z
g (f (inr d)) = inr d
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) pr1: X pr2: Z
g (f (inr (pr1,, pr2))) = inr (pr1,, pr2)
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a
∏ a : X × Y ⨿ Z, f (g a) = a
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a efg: ∏ a : X × Y ⨿ Z, f (g a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a a: X × Y ⨿ Z
f (g a) = a
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a efg: ∏ a : X × Y ⨿ Z, f (g a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a t: X x: Y ⨿ Z
f (g (t,, x)) = t,, x
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a efg: ∏ a : X × Y ⨿ Z, f (g a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a t: X a: Y
f (g (t,, inl a)) = t,, inl a
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a t: X b: Z
f (g (t,, inr b)) = t,, inr b
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a efg: ∏ a : X × Y ⨿ Z, f (g a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a t: X b: Z
f (g (t,, inr b)) = t,, inr b
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a efg: ∏ a : X × Y ⨿ Z, f (g a) = a
isweq f
X, Y, Z: UU f:= rdistrtoprod X Y Z: (X × Y) ⨿ (X × Z) → X × Y ⨿ Z g:= rdistrtocoprod X Y Z: X × Y ⨿ Z → (X × Y) ⨿ (X × Z) egf: ∏ a : (X × Y) ⨿ (X × Z), g (f a) = a efg: ∏ a : X × Y ⨿ Z, f (g a) = a
isweq f
apply (isweq_iso f g egf efg).Defined.Definitionweqrdistrtoprod (XYZ : UU) := make_weq _ (isweqrdistrtoprod X Y Z).
X, Y, Z: UU
isweq (rdistrtocoprod X Y Z)
X, Y, Z: UU
isweq (rdistrtocoprod X Y Z)
X, Y, Z: UU
isweq (rdistrtocoprod X Y Z)
apply (isweqinvmap (weqrdistrtoprod X Y Z)).Defined.Definitionweqrdistrtocoprod (XYZ : UU)
:= make_weq _ (isweqrdistrtocoprod X Y Z).(** *** Total space of a family over a coproduct *)
X, Y: UU P: X ⨿ Y → UU xyp: ∑ y, P y
(∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU xyp: ∑ y, P y
(∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU xyp: ∑ y, P y
(∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU xyp: ∑ y, P y PX:= λx : X, P (inl x): X → UU
(∑ y, PX y) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU xyp: ∑ y, P y PX:= λx : X, P (inl x): X → UU PY:= λy : Y, P (inr y): Y → UU
(∑ y, PX y) ⨿ (∑ y, PY y)
X, Y: UU P: X ⨿ Y → UU xy: X ⨿ Y p: P xy PX:= λx : X, P (inl x): X → UU PY:= λy : Y, P (inr y): Y → UU
(∑ y, PX y) ⨿ (∑ y, PY y)
X, Y: UU P: X ⨿ Y → UU x: X p: P (inl x) PX:= λx : X, P (inl x): X → UU PY:= λy : Y, P (inr y): Y → UU
(∑ y, PX y) ⨿ (∑ y, PY y)
X, Y: UU P: X ⨿ Y → UU y: Y p: P (inr y) PX:= λx : X, P (inl x): X → UU PY:= λy : Y, P (inr y): Y → UU
(∑ y, PX y) ⨿ (∑ y, PY y)
X, Y: UU P: X ⨿ Y → UU y: Y p: P (inr y) PX:= λx : X, P (inl x): X → UU PY:= λy : Y, P (inr y): Y → UU
(∑ y, PX y) ⨿ (∑ y, PY y)
apply (ii2 (tpair PY y p)).Defined.
X, Y: UU P: X ⨿ Y → UU xpyp: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
∑ y, P y
X, Y: UU P: X ⨿ Y → UU xpyp: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
∑ y, P y
X, Y: UU P: X ⨿ Y → UU xpyp: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
∑ y, P y
X, Y: UU P: X ⨿ Y → UU xp: ∑ x : X, P (inl x)
∑ y, P y
X, Y: UU P: X ⨿ Y → UU yp: ∑ y : Y, P (inr y)
∑ y, P y
X, Y: UU P: X ⨿ Y → UU x: X p: P (inl x)
∑ y, P y
X, Y: UU P: X ⨿ Y → UU yp: ∑ y : Y, P (inr y)
∑ y, P y
X, Y: UU P: X ⨿ Y → UU yp: ∑ y : Y, P (inr y)
∑ y, P y
X, Y: UU P: X ⨿ Y → UU y: Y p: P (inr y)
∑ y, P y
apply (tpair P (ii2 y) p).Defined.
X, Y: UU P: X ⨿ Y → UU
(∑ xy : X ⨿ Y, P xy)
≃ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU
(∑ xy : X ⨿ Y, P xy)
≃ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU
(∑ xy : X ⨿ Y, P xy)
≃ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
(∑ xy : X ⨿ Y, P xy)
≃ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y
(∑ xy : X ⨿ Y, P xy)
≃ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y
isweq f
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y
∏ a : ∑ y, P y, g (f a) = a
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
isweq f
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y
∏ a : ∑ y, P y, g (f a) = a
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y a: ∑ y, P y
g (f a) = a
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y xy: X ⨿ Y p: P xy
g (f (xy,, p)) = xy,, p
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y x: X p: P (inl x)
g (f (inl x,, p)) = inl x,, p
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y y: Y p: P (inr y)
g (f (inr y,, p)) = inr y,, p
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y x: X p: P (inl x)
inl x,, p = inl x,, p
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y y: Y p: P (inr y)
g (f (inr y,, p)) = inr y,, p
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y y: Y p: P (inr y)
g (f (inr y,, p)) = inr y,, p
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y y: Y p: P (inr y)
inr y,, p = inr y,, p
apply idpath.
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
isweq f
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
∏ a : (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)),
f (g a) = a
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a efg: ∏ a : (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)), f (g a) = a
isweq f
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
∏ a : (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)),
f (g a) = a
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a a: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
f (g a) = a
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a xp: ∑ x : X, P (inl x)
f (g (inl xp)) = inl xp
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a yp: ∑ y : Y, P (inr y)
f (g (inr yp)) = inr yp
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a x: X p: P (inl x)
f (g (inl (x,, p))) = inl (x,, p)
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a yp: ∑ y : Y, P (inr y)
f (g (inr yp)) = inr yp
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a x: X p: P (inl x)
inl (x,, p) = inl (x,, p)
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a yp: ∑ y : Y, P (inr y)
f (g (inr yp)) = inr yp
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a yp: ∑ y : Y, P (inr y)
f (g (inr yp)) = inr yp
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a y: Y p: P (inr y)
f (g (inr (y,, p))) = inr (y,, p)
apply idpath.
X, Y: UU P: X ⨿ Y → UU f:= fromtotal2overcoprod P: (∑ y, P y)
→ (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)) g:= tototal2overcoprod P: (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y))
→ ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a efg: ∏ a : (∑ x : X, P (inl x)) ⨿ (∑ y : Y, P (inr y)), f (g a) = a
isweq f
apply (isweq_iso _ _ egf efg).Defined.(** *** Pairwise sum of functions, coproduct associativity and commutativity *)Definitionsumofmaps {XYZ : UU} (fx : X -> Z)(fy : Y -> Z) :
(X ⨿ Y) -> Z := λxy : _, match xy with ii1 x => fx x | ii2 y => fy y end.
X, Y, Z: UU
X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z)
X, Y, Z: UU
X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z)
X, Y, Z: UU X0: X ⨿ Y ⨿ Z
X ⨿ (Y ⨿ Z)
X, Y, Z: UU c: X ⨿ Y
X ⨿ (Y ⨿ Z)
X, Y, Z: UU z: Z
X ⨿ (Y ⨿ Z)
X, Y, Z: UU x: X
X ⨿ (Y ⨿ Z)
X, Y, Z: UU y: Y
X ⨿ (Y ⨿ Z)
X, Y, Z: UU z: Z
X ⨿ (Y ⨿ Z)
X, Y, Z: UU y: Y
X ⨿ (Y ⨿ Z)
X, Y, Z: UU z: Z
X ⨿ (Y ⨿ Z)
X, Y, Z: UU z: Z
X ⨿ (Y ⨿ Z)
apply (ii2 (ii2 z)).Defined.
X, Y, Z: UU
X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z
X, Y, Z: UU
X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z
X, Y, Z: UU X0: X ⨿ (Y ⨿ Z)
X ⨿ Y ⨿ Z
X, Y, Z: UU x: X
X ⨿ Y ⨿ Z
X, Y, Z: UU c: Y ⨿ Z
X ⨿ Y ⨿ Z
X, Y, Z: UU c: Y ⨿ Z
X ⨿ Y ⨿ Z
X, Y, Z: UU y: Y
X ⨿ Y ⨿ Z
X, Y, Z: UU z: Z
X ⨿ Y ⨿ Z
X, Y, Z: UU z: Z
X ⨿ Y ⨿ Z
apply (ii2 z).Defined.
X, Y, Z, T: UU f: X → T g: Y → T h: Z → T
sumofmaps (sumofmaps f g) h ∘ coprodasstol X Y Z ~
sumofmaps f (sumofmaps g h)
X, Y, Z, T: UU f: X → T g: Y → T h: Z → T
sumofmaps (sumofmaps f g) h ∘ coprodasstol X Y Z ~
sumofmaps f (sumofmaps g h)
X, Y, Z, T: UU f: X → T g: Y → T h: Z → T
sumofmaps (sumofmaps f g) h ∘ coprodasstol X Y Z ~
sumofmaps f (sumofmaps g h)
intros [x|[y|z]]; apply idpath.Defined.
X, Y, Z, T: UU f: X → T g: Y → T h: Z → T
sumofmaps f (sumofmaps g h) ∘ coprodasstor X Y Z ~
sumofmaps (sumofmaps f g) h
X, Y, Z, T: UU f: X → T g: Y → T h: Z → T
sumofmaps f (sumofmaps g h) ∘ coprodasstor X Y Z ~
sumofmaps (sumofmaps f g) h
X, Y, Z, T: UU f: X → T g: Y → T h: Z → T
sumofmaps f (sumofmaps g h) ∘ coprodasstor X Y Z ~
sumofmaps (sumofmaps f g) h
intros [[x|y]|z]; apply idpath.Defined.
X, Y, Z: UU
isweq (coprodasstor X Y Z)
X, Y, Z: UU
isweq (coprodasstor X Y Z)
X, Y, Z: UU
isweq (coprodasstor X Y Z)
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z)
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z
∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z xyz: X ⨿ Y ⨿ Z
g (f xyz) = xyz
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z c: X ⨿ Y
g (f (inl c)) = inl c
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z z: Z
g (f (inr z)) = inr z
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z a: X
g (f (inl (inl a))) = inl (inl a)
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z b: Y
g (f (inl (inr b))) = inl (inr b)
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z z: Z
g (f (inr z)) = inr z
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z b: Y
g (f (inl (inr b))) = inl (inr b)
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z z: Z
g (f (inr z)) = inr z
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z z: Z
g (f (inr z)) = inr z
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz
∏ xyz : X ⨿ (Y ⨿ Z), f (g xyz) = xyz
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz efg: ∏ xyz : X ⨿ (Y ⨿ Z), f (g xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz xyz: X ⨿ (Y ⨿ Z)
f (g xyz) = xyz
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz efg: ∏ xyz : X ⨿ (Y ⨿ Z), f (g xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz x: X
f (g (inl x)) = inl x
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz c: Y ⨿ Z
f (g (inr c)) = inr c
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz efg: ∏ xyz : X ⨿ (Y ⨿ Z), f (g xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz c: Y ⨿ Z
f (g (inr c)) = inr c
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz efg: ∏ xyz : X ⨿ (Y ⨿ Z), f (g xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz a: Y
f (g (inr (inl a))) = inr (inl a)
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz b: Z
f (g (inr (inr b))) = inr (inr b)
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz efg: ∏ xyz : X ⨿ (Y ⨿ Z), f (g xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz b: Z
f (g (inr (inr b))) = inr (inr b)
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz efg: ∏ xyz : X ⨿ (Y ⨿ Z), f (g xyz) = xyz
isweq f
X, Y, Z: UU f:= coprodasstor X Y Z: X ⨿ Y ⨿ Z → X ⨿ (Y ⨿ Z) g:= coprodasstol X Y Z: X ⨿ (Y ⨿ Z) → X ⨿ Y ⨿ Z egf: ∏ xyz : X ⨿ Y ⨿ Z, g (f xyz) = xyz efg: ∏ xyz : X ⨿ (Y ⨿ Z), f (g xyz) = xyz
isweq f
apply (isweq_iso f g egf efg).Defined.Definitionweqcoprodasstor (XYZ : UU) := make_weq _ (isweqcoprodasstor X Y Z).
X, Y, Z: UU
isweq (coprodasstol X Y Z)
X, Y, Z: UU
isweq (coprodasstol X Y Z)
X, Y, Z: UU
isweq (coprodasstol X Y Z)
apply (isweqinvmap (weqcoprodasstor X Y Z)).Defined.Definitionweqcoprodasstol (XYZ : UU) := make_weq _ (isweqcoprodasstol X Y Z).Definitioncoprodcomm (XY : UU) : X ⨿ Y -> Y ⨿ X
:= λxy : _, match xy with ii1 x => ii2 x | ii2 y => ii1 y end.
X, Y: UU
isweq (coprodcomm X Y)
X, Y: UU
isweq (coprodcomm X Y)
X, Y: UU
isweq (coprodcomm X Y)
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y
∏ xy : X ⨿ Y, g (f xy) = xy
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y xy: X ⨿ Y
g (f xy) = xy
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y a: X
g (f (inl a)) = inl a
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y b: Y
g (f (inr b)) = inr b
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y b: Y
g (f (inr b)) = inr b
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
∏ yx : Y ⨿ X, f (g yx) = yx
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ yx : Y ⨿ X, f (g yx) = yx
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy yx: Y ⨿ X
f (g yx) = yx
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ yx : Y ⨿ X, f (g yx) = yx
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy a: Y
f (g (inl a)) = inl a
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy b: X
f (g (inr b)) = inr b
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ yx : Y ⨿ X, f (g yx) = yx
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy b: X
f (g (inr b)) = inr b
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ yx : Y ⨿ X, f (g yx) = yx
isweq f
X, Y: UU f:= coprodcomm X Y: X ⨿ Y → Y ⨿ X g:= coprodcomm Y X: Y ⨿ X → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ yx : Y ⨿ X, f (g yx) = yx
isweq f
apply (isweq_iso f g egf efg).Defined.Definitionweqcoprodcomm (XY : UU) := make_weq _ (isweqcoprodcomm X Y).(** *** Coproduct with a "negative" type *)
X, Y: UU nf: Y → ∅
isweq inl
X, Y: UU nf: Y → ∅
isweq inl
X, Y: UU nf: Y → ∅
isweq inl
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X
∏ x : X, g (f x) = x
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X x: X
g (f x) = x
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x
∏ xy : X ⨿ Y, f (g xy) = xy
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x xy: X ⨿ Y
f (g xy) = xy
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x x: X
f (g (inl x)) = inl x
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x y: Y
f (g (inr y)) = inr y
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x y: Y
f (g (inr y)) = inr y
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
X, Y: UU nf: Y → ∅ f:= inl: X → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => x
| inr y => fromempty (nf y)
end: X ⨿ Y → X egf: ∏ x : X, g (f x) = x efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
apply (isweq_iso f g egf efg).Defined.Definitionweqii1withneg (X : UU) {Y : UU} (nf : ¬ Y)
:= make_weq _ (isweqii1withneg X nf).
X, Y: UU nf: X → ∅
isweq inr
X, Y: UU nf: X → ∅
isweq inr
X, Y: UU nf: X → ∅
isweq inr
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y
∏ y : Y, g (f y) = y
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y y: Y
g (f y) = y
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y
∏ xy : X ⨿ Y, f (g xy) = xy
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y xy: X ⨿ Y
f (g xy) = xy
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y x: X
f (g (inl x)) = inl x
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y y: Y
f (g (inr y)) = inr y
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y y: Y
f (g (inr y)) = inr y
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
X, Y: UU nf: X → ∅ f:= inr: Y → X ⨿ Y g:= λxy : X ⨿ Y,
match xy with
| inl x => fromempty (nf x)
| inr y => y
end: X ⨿ Y → Y egf: ∏ y : Y, g (f y) = y efg: ∏ xy : X ⨿ Y, f (g xy) = xy
isweq f
apply (isweq_iso f g egf efg).Defined.Definitionweqii2withneg {X : UU} (Y : UU) (nf : ¬ X)
:= make_weq _ (isweqii2withneg Y nf).(** *** Coproduct of two functions *)Definitioncoprodf {XYX'Y' : UU} (f : X -> X') (g : Y-> Y') : X ⨿ Y -> X' ⨿ Y'
:= λxy: X ⨿ Y,
match xy with
| ii1 x => ii1 (f x)
| ii2 y => ii2 (g y)
end.
X, Y, X': UU
(X → X') → X ⨿ Y → X' ⨿ Y
X, Y, X': UU
(X → X') → X ⨿ Y → X' ⨿ Y
X, Y, X': UU f: X → X'
X ⨿ Y → X' ⨿ Y
exact (coprodf f (idfun Y)).Defined.
X, Y, Y': UU
(Y → Y') → X ⨿ Y → X ⨿ Y'
X, Y, Y': UU
(Y → Y') → X ⨿ Y → X ⨿ Y'
X, Y, Y': UU g: Y → Y'
X ⨿ Y → X ⨿ Y'
exact (coprodf (idfun X) g).Defined.
X, X', Y, Y', Z, Z': UU f: X → Y f': X' → Y' g: Y → Z g': Y' → Z'
coprodf g g' ∘ coprodf f f' ~
coprodf (g ∘ f) (g' ∘ f')
X, X', Y, Y', Z, Z': UU f: X → Y f': X' → Y' g: Y → Z g': Y' → Z'
coprodf g g' ∘ coprodf f f' ~
coprodf (g ∘ f) (g' ∘ f')
X, X', Y, Y', Z, Z': UU f: X → Y f': X' → Y' g: Y → Z g': Y' → Z'
coprodf g g' ∘ coprodf f f' ~
coprodf (g ∘ f) (g' ∘ f')
X, X', Y, Y', Z, Z': UU f: X → Y f': X' → Y' g: Y → Z g': Y' → Z' xx': X ⨿ X'
(coprodf g g' ∘ coprodf f f') xx' =
coprodf (g ∘ f) (g' ∘ f') xx'
X, X', Y, Y', Z, Z': UU f: X → Y f': X' → Y' g: Y → Z g': Y' → Z' x: X
(coprodf g g' ∘ coprodf f f') (inl x) =
coprodf (g ∘ f) (g' ∘ f') (inl x)
X, X', Y, Y', Z, Z': UU f: X → Y f': X' → Y' g: Y → Z g': Y' → Z' x': X'
(coprodf g g' ∘ coprodf f f') (inr x') =
coprodf (g ∘ f) (g' ∘ f') (inr x')
X, X', Y, Y', Z, Z': UU f: X → Y f': X' → Y' g: Y → Z g': Y' → Z' x': X'
(coprodf g g' ∘ coprodf f f') (inr x') =
coprodf (g ∘ f) (g' ∘ f') (inr x')
apply idpath.Defined.Definitionhomotcoprodfhomot {XX'YY' : UU} (fg : X -> Y)
(f'g' : X' -> Y') (h : homot f g) (h' : homot f' g') :
homot (coprodf f f') (coprodf g g')
:= λxx' : _, match xx' with
| ii1 x => maponpaths (@ii1 _ _) (h x)
| ii2 x' => maponpaths (@ii2 _ _) (h' x')
end.
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y'
isweq (coprodf w w')
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y'
isweq (coprodf w w')
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y'
isweq (coprodf w w')
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X
isweq (coprodf w w')
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y
isweq (coprodf w w')
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y'
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y
∏ xy : X ⨿ Y, gg (ff xy) = xy
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y xy: X ⨿ Y
gg (ff xy) = xy
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y x: X
gg (ff (inl x)) = inl x
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y y: Y
gg (ff (inr y)) = inr y
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y x: X
inl (finv (w x)) = inl x
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y y: Y
gg (ff (inr y)) = inr y
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y y: Y
gg (ff (inr y)) = inr y
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy
∏ xy' : X' ⨿ Y', ff (gg xy') = xy'
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy efg: ∏ xy' : X' ⨿ Y', ff (gg xy') = xy'
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy xy': X' ⨿ Y'
ff (gg xy') = xy'
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy efg: ∏ xy' : X' ⨿ Y', ff (gg xy') = xy'
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy x: X'
ff (gg (inl x)) = inl x
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy y: Y'
ff (gg (inr y)) = inr y
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy efg: ∏ xy' : X' ⨿ Y', ff (gg xy') = xy'
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy x: X'
inl (w (finv x)) = inl x
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy y: Y'
ff (gg (inr y)) = inr y
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy efg: ∏ xy' : X' ⨿ Y', ff (gg xy') = xy'
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy y: Y'
ff (gg (inr y)) = inr y
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy efg: ∏ xy' : X' ⨿ Y', ff (gg xy') = xy'
isweq ff
X, Y, X', Y': UU w: X ≃ X' w': Y ≃ Y' finv:= invmap w: X' → X ginv:= invmap w': Y' → Y ff:= coprodf w w': X ⨿ Y → X' ⨿ Y' gg:= coprodf finv ginv: X' ⨿ Y' → X ⨿ Y egf: ∏ xy : X ⨿ Y, gg (ff xy) = xy efg: ∏ xy' : X' ⨿ Y', ff (gg xy') = xy'
isweq ff
apply (isweq_iso ff gg egf efg).Defined.
X, Y, X', Y': UU
X ≃ X' → Y ≃ Y' → X ⨿ Y ≃ X' ⨿ Y'
X, Y, X', Y': UU
X ≃ X' → Y ≃ Y' → X ⨿ Y ≃ X' ⨿ Y'
X, Y, X', Y': UU w1: X ≃ X' w2: Y ≃ Y'
X ⨿ Y ≃ X' ⨿ Y'
exact (make_weq _ (isweqcoprodf w1 w2)).Defined.
X, Y, X': UU
X ≃ X' → X ⨿ Y ≃ X' ⨿ Y
X, Y, X': UU
X ≃ X' → X ⨿ Y ≃ X' ⨿ Y
X, Y, X': UU w: X ≃ X'
X ⨿ Y ≃ X' ⨿ Y
exact (weqcoprodf w (idweq Y)).Defined.
X, Y, Y': UU
Y ≃ Y' → X ⨿ Y ≃ X ⨿ Y'
X, Y, Y': UU
Y ≃ Y' → X ⨿ Y ≃ X ⨿ Y'
X, Y, Y': UU w: Y ≃ Y'
X ⨿ Y ≃ X ⨿ Y'
exact (weqcoprodf (idweq X) w).Defined.(** *** The [ equality_cases ] construction and four applications to [ ii1 ] and [ ii2 ] *)(* Added by D. Grayson, Nov. 2015 *)
exact (maponpaths (@ii2 P Q) e).Defined.(* the same proof proves 4 lemmas: *)
P, Q: UU p, p': P
inl p = inl p' → p = p'
P, Q: UU p, p': P
inl p = inl p' → p = p'
exact equality_by_case.Defined.
P, Q: UU q, q': Q
inr q = inr q' → q = q'
P, Q: UU q, q': Q
inr q = inr q' → q = q'
exact equality_by_case.Defined.
X, Y: UU x: X y: Y
inl x != inr y
X, Y: UU x: X y: Y
inl x != inr y
exact equality_by_case.Defined.
X, Y: UU x: X y: Y
inr y != inl x
X, Y: UU x: X y: Y
inr y != inl x
exact equality_by_case.Defined.(** *** Bool as coproduct *)
unit ⨿ unit ≃ bool
unit ⨿ unit ≃ bool
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool
unit ⨿ unit ≃ bool
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit
∏ xx : unit ⨿ unit, g (f xx) = xx
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit xx: unit ⨿ unit
g (f xx) = xx
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit u: unit
g (f (inl u)) = inl u
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit u: unit
g (f (inr u)) = inr u
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit
g (f (inl tt)) = inl tt
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit u: unit
g (f (inr u)) = inr u
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit u: unit
g (f (inr u)) = inr u
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit
g (f (inr tt)) = inr tt
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
∏ t : bool, f (g t) = t
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx efg: ∏ t : bool, f (g t) = t
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
f (g true) = true
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
f (g false) = false
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx efg: ∏ t : bool, f (g t) = t
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx
f (g false) = false
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx efg: ∏ t : bool, f (g t) = t
isweq f
f:= λxx : unit ⨿ unit,
match xx with
| inl _ => true
| inr _ => false
end: unit ⨿ unit → bool g:= λt : bool,
match t with
| true => inl tt
| false => inr tt
end: bool → unit ⨿ unit egf: ∏ xx : unit ⨿ unit, g (f xx) = xx efg: ∏ t : bool, f (g t) = t
isweq f
apply (isweq_iso f g egf efg).Defined.(** *** Pairwise coproducts as dependent sums of families over [ bool ] *)Definitioncoprodtobool {XY : UU} (xy : X ⨿ Y) : bool
:= match xy with
| ii1 x => true
| ii2 y => false
end.Definitionboolsumfun (XY : UU) : bool -> UU
:= λt : _, match t with
| true => X
| false => Y
end.Definitioncoprodtoboolsum (XY : UU) : X ⨿ Y -> total2 (boolsumfun X Y)
:= λxy : _, match xy with
| ii1 x => tpair (boolsumfun X Y) true x
| ii2 y => tpair (boolsumfun X Y) false y
end.Definitionboolsumtocoprod (XY : UU): (total2 (boolsumfun X Y)) -> X ⨿ Y
:= λxy, match xy with
| tpair _ true x => ii1 x
| tpair _ false y => ii2 y
end.
X, Y: UU
isweq (coprodtoboolsum X Y)
X, Y: UU
isweq (coprodtoboolsum X Y)
X, Y: UU
isweq (coprodtoboolsum X Y)
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y
∏ xy : X ⨿ Y, g (f xy) = xy
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y a: X
g (f (inl a)) = inl a
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y b: Y
g (f (inr b)) = inr b
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y b: Y
g (f (inr b)) = inr b
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy
∏ xy : ∑ y, boolsumfun X Y y, f (g xy) = xy
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ xy : ∑ y, boolsumfun X Y y, f (g xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy xy: ∑ y, boolsumfun X Y y
f (g xy) = xy
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ xy : ∑ y, boolsumfun X Y y, f (g xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy t: bool x: boolsumfun X Y t
f (g (t,, x)) = t,, x
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ xy : ∑ y, boolsumfun X Y y, f (g xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy x: boolsumfun X Y true
f (g (true,, x)) = true,, x
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy x: boolsumfun X Y false
f (g (false,, x)) = false,, x
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ xy : ∑ y, boolsumfun X Y y, f (g xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy x: boolsumfun X Y false
f (g (false,, x)) = false,, x
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ xy : ∑ y, boolsumfun X Y y, f (g xy) = xy
isweq f
X, Y: UU f:= coprodtoboolsum X Y: X ⨿ Y → ∑ y, boolsumfun X Y y g:= boolsumtocoprod X Y: (∑ y, boolsumfun X Y y) → X ⨿ Y egf: ∏ xy : X ⨿ Y, g (f xy) = xy efg: ∏ xy : ∑ y, boolsumfun X Y y, f (g xy) = xy
isweq f
apply (isweq_iso f g egf efg).Defined.Definitionweqcoprodtoboolsum (XY : UU) := make_weq _ (isweqcoprodtoboolsum X Y).
X, Y: UU
isweq (boolsumtocoprod X Y)
X, Y: UU
isweq (boolsumtocoprod X Y)
X, Y: UU
isweq (boolsumtocoprod X Y)
apply (isweqinvmap (weqcoprodtoboolsum X Y)).Defined.Definitionweqboolsumtocoprod (XY : UU) := make_weq _ (isweqboolsumtocoprod X Y).(** *** Splitting of [ X ] into a coproduct defined by a function [ X -> Y ⨿ Z ] *)
X, Y, Z: UU f: X → Y ⨿ Z
X
≃ (∑ y : Y, hfiber f (inl y))
⨿ (∑ z : Z, hfiber f (inr z))
X, Y, Z: UU f: X → Y ⨿ Z
X
≃ (∑ y : Y, hfiber f (inl y))
⨿ (∑ z : Z, hfiber f (inr z))
X, Y, Z: UU f: X → Y ⨿ Z
X
≃ (∑ y : Y, hfiber f (inl y))
⨿ (∑ z : Z, hfiber f (inr z))
X, Y, Z: UU f: X → Y ⨿ Z w1:= weqtococonusf f: X ≃ coconusf f
X
≃ (∑ y : Y, hfiber f (inl y))
⨿ (∑ z : Z, hfiber f (inr z))
X, Y, Z: UU f: X → Y ⨿ Z w1:= weqtococonusf f: X ≃ coconusf f w2:= weqtotal2overcoprod (λyz : Y ⨿ Z, hfiber f yz): (∑ xy : Y ⨿ Z, (λyz : Y ⨿ Z, hfiber f yz) xy)
≃ (∑ x : Y, (λyz : Y ⨿ Z, hfiber f yz) (inl x))
⨿ (∑ y : Z, (λyz : Y ⨿ Z, hfiber f yz) (inr y))
X
≃ (∑ y : Y, hfiber f (inl y))
⨿ (∑ z : Z, hfiber f (inr z))
apply (weqcomp w1 w2).Defined.(** *** Some properties of [ bool ] *)
x: bool
(x = true) ⨿ (x = false)
x: bool
(x = true) ⨿ (x = false)
(true = true) ⨿ (true = false)
(false = true) ⨿ (false = false)
(false = true) ⨿ (false = false)
apply (ii2 (idpath _)).Defined.
bool → UU
bool → UU
b: bool
UU
UU
UU
UU
exact unit.
UU
UU
exact empty.}Defined.
true = false → ∅
true = false → ∅
X: true = false
∅
apply (transportf bool_to_type X tt).Defined.
false = true → ∅
false = true → ∅
X: false = true
∅
apply (transportb bool_to_type X tt).Defined.
x: bool
x = true → x != false
x: bool
x = true → x != false
x: bool e: x = true
x != false
x: bool e: x = true
true != false
x: bool e: x = true
true = false → ∅
apply nopathstruetofalse.Defined.
x: bool
x = false → x != true
x: bool
x = false → x != true
x: bool e: x = false
x != true
x: bool e: x = false
false != true
x: bool e: x = false
false = true → ∅
apply nopathsfalsetotrue.Defined.
x: bool
x != true → x = false
x: bool
x != true → x = false
x: bool ne: x != true
x = false
x: bool ne: x != true t: x = true
x = false
x: bool ne: x != true f: x = false
x = false
x: bool ne: x != true f: x = false
x = false
apply f.Defined.
x: bool
x != false → x = true
x: bool
x != false → x = true
x: bool ne: x != false
x = true
x: bool ne: x != false t: x = true
x = true
x: bool ne: x != false f: x = false
x = true
x: bool ne: x != false f: x = false
x = true
induction (ne f).Defined.(** *** Fibrations with only one non-empty fiberTheorem saying that if a fibration has only one non-empty fiber then the totalspace is weakly equivalent to this fiber. *)(* The current proof added by P. L. Lumsdaine, 2016 *)
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x'
isweq (λp : P x, x,, p)
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x'
isweq (λp : P x, x,, p)
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x'
isweq (λp : P x, x,, p)
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x
∏ x' : X, (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= ?Goal: ∏ x' : X, (x = x') ⨿ ¬ P x'
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x
∏ x' : X, (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X
(x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X
x = x → (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X
¬ P x → (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X
x = x → (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X x0: x = x
(x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X x0: x = x
x = x' → (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X x0: x = x
¬ P x' → (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X x0: x = x
x = x' → (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X x0: x = x ee: x = x'
(x = x') ⨿ ¬ P x'
apply ii1; exact (pathscomp0 (pathsinv0 x0) ee).
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X x0: x = x
¬ P x' → (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X x0: x = x phi: ¬ P x'
(x = x') ⨿ ¬ P x'
apply ii2, phi.
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X
¬ P x → (x = x') ⨿ ¬ P x'
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x x': X phi: ¬ P x
(x = x') ⨿ ¬ P x'
exact (c x').
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x'
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x
∏ pp : ∑ y, P y, f (g pp) = pp
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x pp: ∑ y, P y
f (g pp) = pp
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t
f (g (t,, x0)) = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t cnewt:= cnew t: (x = t) ⨿ ¬ P t
f (g (t,, x0)) = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t cnewt:= cnew t: (x = t) ⨿ ¬ P t
f
match cnew (pr1 (t,, x0)) with
| inl e => transportb P e (pr2 (t,, x0))
| inr phi => fromempty (phi (pr2 (t,, x0)))
end = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t cnewt:= cnew t: (x = t) ⨿ ¬ P t
x,,
match cnew (pr1 (t,, x0)) with
| inl e => transportb P e (pr2 (t,, x0))
| inr phi => fromempty (phi (pr2 (t,, x0)))
end = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t cnewt:= cnew t: (x = t) ⨿ ¬ P t
x,,
match cnew t with
| inl e => transportb P e x0
| inr phi => fromempty (phi x0)
end = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t cnewt:= cnew t: (x = t) ⨿ ¬ P t
x,,
match cnewt with
| inl e => transportb P e x0
| inr phi => fromempty (phi x0)
end = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t x1: x = t
x,, transportb P x1 x0 = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t y: ¬ P t
x,, fromempty (y x0) = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x t: X x0: P t y: ¬ P t
x,, fromempty (y x0) = t,, x0
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= cnew x: (x = x) ⨿ ¬ P x
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= cnew x: (x = x) ⨿ ¬ P x
cnew x = cnewx
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= cnew x: (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= cnew x: (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(c x)) (λ_ : ¬ P x, c x) cx: (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y cx:= c x: (x = x) ⨿ ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x') cx: ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi) cx)
(λ_ : ¬ P x, cx) cx: (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
cnewx = inl (idpath x)
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x)
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x)
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x)
∏ p : P x, g (f p) = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x
g (f p) = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x
g (f p) = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x
match cnew (pr1 (f p)) with
| inl e => transportb P e (pr2 (f p))
| inr phi => fromempty (phi (pr2 (f p)))
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x
match cnew (pr1 (x,, p)) with
| inl e => transportb P e (pr2 (x,, p))
| inr phi => fromempty (phi (pr2 (x,, p)))
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x
ff cnewx = ff (inl (idpath x))
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x))
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x))
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x))
ff (inl (idpath x)) = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x)) eee: ff (inl (idpath x)) = p
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x)) eee: ff (inl (idpath x)) = p
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x)) eee: ff (inl (idpath x)) = p
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x)) eee: ff (inl (idpath x)) = p
ff (cnew x) = ff cnewx
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x)) eee: ff (inl (idpath x)) = p e2: ff (cnew x) = ff cnewx
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) p: P x ff:= λcc : (x = x) ⨿ (P x → ∅),
match cc with
| inl e0 => transportb P e0 p
| inr phi => fromempty (phi p)
end: (x = x) ⨿ (P x → ∅) → P x ee: ff cnewx = ff (inl (idpath x)) eee: ff (inl (idpath x)) = p e2: ff (cnew x) = ff cnewx
match cnew x with
| inl e => transportb P e p
| inr phi => fromempty (phi p)
end = p
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y x0: x = x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inl x0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx1 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x1 @ ee))
(λphi : ¬ P x, inr phi)
(inl x0)) (λ_ : ¬ P x, inl x0)
(inl x0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx e: cnewx = inl (idpath x) egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
isweq f
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx
∏ y : ∑ y, P y, iscontr (hfiber f y)
X: UU P: X → UU x: X c: ∏ x' : X, (x = x') ⨿ ¬ P x' f:= λp : P x, x,, p: P x → ∑ y, P y e0: ¬ P x cnew:= λx' : X,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x') ⨿ ¬ P x')
(λx0 : x = x,
coprod_rect
(λ_ : (x = x') ⨿ ¬ P x', (x = x') ⨿ ¬ P x')
(λee : x = x', inl (! x0 @ ee))
(λphi : ¬ P x', inr phi)
(c x')) (λ_ : ¬ P x, c x')
(inr e0): ∏ x' : X, (x = x') ⨿ ¬ P x' g:= λpp : ∑ y, P y,
match cnew (pr1 pp) with
| inl e => transportb P e (pr2 pp)
| inr phi => fromempty (phi (pr2 pp))
end: (∑ y, P y) → P x efg: ∏ pp : ∑ y, P y, f (g pp) = pp cnewx:= coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λx0 : x = x,
coprod_rect
(λ_ : (x = x) ⨿ ¬ P x, (x = x) ⨿ ¬ P x)
(λee : x = x, inl (! x0 @ ee))
(λphi : ¬ P x, inr phi)
(inr e0)) (λ_ : ¬ P x, inr e0)
(inr e0): (x = x) ⨿ ¬ P x e1: cnew x = cnewx y0: ∑ y, P y
iscontr (hfiber f y0)
induction (e0 (g y0)).Defined.(** ** Basics about fibration sequences. *)(** *** Fibrations sequences and their first "left shifts"The group of constructions related to fibration sequences forms one of the mostimportant computational toolboxes of homotopy theory.Given a pair of functions [ (f : X -> Y) (g : Y -> Z) ] and a point [ z : Z ] ,a structure of the complex on such a triple is a homotopy from the composition[ funcomp f g ] to the constant function [ X -> Z ] corresponding to [ z ] i.e.a term [ ez : ∏ x : X, (g (f x)) = z ]. Specifing such a structure isessentially equivalent to specifing a structure of the form[ ezmap : X -> hfiber g z ]. The mapping in one direction is given in thedefinition of [ ezmap ] below. The mapping in another is given by[ f := λ x : X, pr1 (ezmap x) ] and [ ez := λ x : X, pr2 (ezmap x) ].A complex is called a fibration sequence if [ ezmap ] is a weak equivalence.Correspondingly, the structure of a fibration sequence on [ f g z ] is a pair[ (ez , is) ] where [ is : isweq (ezmap f g z ez) ]. For a fibration sequence[ f g z fs ] where [ fs : fibseqstr f g z ] and any [ y : Y ] there is defineda function [ diff1 : (g y) = z -> X ] and a structure of the fibrationsequence [ fibseqdiff1 ] on the triple [ diff1 g y ]. This new fibrationsequence is called the derived fibration sequence of the original one.The first function of the second derived of [ f g z fs ] corresponding to[ (y : Y) (x : X) ] is of the form [ (f x) = y -> (g y) = z ] and it ishomotopic to the function defined by[ e => pathscomp0 (maponpaths g (pathsinv0 e)) (ez x) ]. The first function ofthe third derived of [ f g z fs ] corresponding to[ (y : Y) (x : X) (e : (g y) = z) ] is of the form[ (diff1 e) = x -> (f x) = y ]. Therefore, the third derived of a sequence basedon [ X Y Z ] is based entirely on types = of [ X ], [ Y ] and [ Z ]. When thisconstruction is applied to types of finite h-level (see below) and combined withthe fact that the h-level of a path type is strictly lower than the h-level ofthe ambient type it leads to the possibility of building proofs about types byinduction on h-level.There are three important special cases in which fibration sequences arise:( pr1 - case ) The fibration sequence [ fibseqpr1 P z ] defined by family[ P : Z -> UU ] and a term [ z : Z ]. It is based on the sequence of functions[ (tpair P z : P z -> total2 P) (pr1 : total2 P -> Z) ]. The corresponding[ ezmap ] is defined by an obvious rule and the fact that it is a weakequivalence is proved in [ isweqfibertohfiber ].( g - case ) The fibration sequence [ fibseqg g z ] defined by a function[ g : Y -> Z ] and a term [ z : Z ]. It is based on the sequence of functions[ (hfiberpr1 : hfiber g z -> Y) (g : Y -> Z) ] and the corresponding [ ezmap ]is the function which takes a term [ ye : hfiber ] to[ make_hfiber g (pr1 ye) (pr2 ye) ]. We now have eta-coercion for dependentsums, so it is the identity function,which is sufficient to ensure that it is a weak equivalence. The first derivedof [ fibseqg g z ] corresponding to [ y : Y ] coincides with[ fibseqpr1 (λ y' : Y , (g y') = z) y ].( hf -case ) The fibration sequence of homotopy fibers defined for any pair offunctions [ (f : X -> Y) (g : Y -> Z) ] and any terms[ (z : Z) (ye : hfiber g z) ]. It is based on functions[ hfiberftogf : hfiber f (pr1 ye) -> hfiber (funcomp f g) z ] and[ hfibergftog : hfiber (funcomp f g) z -> hfiber g z ] which are defined below. *)(** *** The structures of a complex and of a fibration sequence on a composable pair of functions *)(** The structure of a complex on a composable pair of functions [ (f : X -> Y) (g : Y -> Z) ] relative to a term [ z : Z ]. *)Definitioncomplxstr {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
:= ∏ x : X, (g (f x)) = z.(** The structure of a fibration sequence on a complex. *)Definitionezmap {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(ez : complxstr f g z) :
X -> hfiber g z := λx : X, make_hfiber g (f x) (ez x).Definitionisfibseq {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(ez : complxstr f g z) := isweq (ezmap f g z ez).Definitionfibseqstr {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
:= ∑ ez : complxstr f g z, isfibseq f g z ez.Definitionmake_fibseqstr {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
:= tpair (λez : complxstr f g z, isfibseq f g z ez).Definitionfibseqstrtocomplxstr {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z) :
fibseqstr f g z -> complxstr f g z
:= @pr1 _ (λez : complxstr f g z, isfibseq f g z ez).Coercionfibseqstrtocomplxstr : fibseqstr >-> complxstr.Definitionezweq {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(fs : fibseqstr f g z) : X ≃ hfiber g z
:= make_weq _ (pr2 fs).(** *** Construction of the derived fibration sequence *)Definitiond1 {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(fs : fibseqstr f g z) (y : Y) : (g y) = z -> X
:= λe : _, invmap (ezweq f g z fs) (make_hfiber g y e).
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y e: g y = z
hfiber f y
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y e: g y = z
hfiber f y
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y e: g y = z
hfiber f y
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y e: g y = z
f (d1 f g z fs y e) = y
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y e: g y = z
f (invmap (ezweq f g z fs) (make_hfiber g y e)) = y
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y e: g y = z
hfiberpr1 g z
(ezweq f g z fs
(invmap (ezweq f g z fs) (make_hfiber g y e))) =
y
apply (maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (make_hfiber g y e))).Defined.Definitioninvezmap1 {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(ez : complxstr f g z) (y : Y) : hfiber f y -> (g y) = z
:= λxe: hfiber f y,
match xe with
tpair _ x e => pathscomp0 (maponpaths g (pathsinv0 e)) (ez x)
end.
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y
isweq (ezmap1 f g z fs y)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y
isweq (ezmap1 f g z fs y)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y
isweq (ezmap1 f g z fs y)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z
∏ e : g y = z, gg (ff e) = e
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z e: g y = z
gg (ff e) = e
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z e: g y = z
matchmatchmatch
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(make_hfiber g y e)) in (_ = a0)
return (a0 = f (d1 f g z fs y e))
with
| idpath _ => idpath (f (d1 f g z fs y e))
endin (_ = a0) return (g y = g a0)
with
| idpath _ => idpath (g y)
endin (_ = a0) return (a0 = z → g y = z)
with
| idpath _ => λe2 : g y = z, e2
end (pr1 fs (d1 f g z fs y e)) = e
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e
∏ xe : hfiber f y, ff (gg xe) = xe
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e xe: hfiber f y
ff (gg xe) = xe
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e x: X e: f x = y
ff (gg (x,, e)) = x,, e
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e
ff (gg (x,, idpath (f x))) = x,, idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e
ff (pr1 fs x) = x,, idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e
ezmap1 f g z fs (f x) (pr1 fs x) = x,, idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e
d1 f g z fs (f x) (pr1 fs x),,
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(make_hfiber g (f x) (pr1 fs x))) =
x,, idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e
invmap (ezweq f g z fs)
(make_hfiber g (f x) (pr1 fs x)),,
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(make_hfiber g (f x) (pr1 fs x))) =
x,, idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e
invmap (ezweq f g z fs) (ezmap f g z fs x),,
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
x,, idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e
pr2
(make_hfiber f
(invmap (ezweq f g z fs) (ezmap f g z fs x))
(maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z fs x)))) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
pr2 (make_hfiber f x (idpath (f x)))
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e2: maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e3:= maponpaths
(λe : ezweq f g z fs
(invmap (ezweq f g z fs)
(ezweq f g z fs x)) =
ezweq f g z fs x,
maponpaths (hfiberpr1 g z) e)
(! homotweqinvweqweq (ezweq f g z fs) x): (λe : ezweq f g z fs
(invmap (ezweq f g z fs)
(ezweq f g z fs x)) =
ezweq f g z fs x,
maponpaths (hfiberpr1 g z) e)
(homotweqinvweq (ezweq f g z fs)
(ezweq f g z fs x)) =
(λe : ezweq f g z fs
(invmap (ezweq f g z fs)
(ezweq f g z fs x)) =
ezweq f g z fs x,
maponpaths (hfiberpr1 g z) e)
(maponpaths (ezweq f g z fs)
(homotinvweqweq (ezweq f g z fs) x))
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e2: maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e3:= maponpaths
(λe : ezmap f g z (pr1 fs)
(invmap (ezweq f g z fs)
(ezmap f g z (pr1 fs) x)) =
ezmap f g z (pr1 fs) x,
maponpaths (hfiberpr1 g z) e)
(! homotweqinvweqweq (ezweq f g z fs) x): maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z (pr1 fs) x)) =
maponpaths (hfiberpr1 g z)
(maponpaths (ezmap f g z (pr1 fs))
(homotinvweqweq (ezweq f g z fs) x))
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e2: maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e3:= maponpaths
(λe : ezmap f g z (pr1 fs)
(invmap (ezweq f g z fs)
(ezmap f g z (pr1 fs) x)) =
ezmap f g z (pr1 fs) x,
maponpaths (hfiberpr1 g z) e)
(! homotweqinvweqweq (ezweq f g z fs) x): maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z (pr1 fs) x)) =
maponpaths (hfiberpr1 g z)
(maponpaths (ezmap f g z (pr1 fs))
(homotinvweqweq (ezweq f g z fs) x)) e4:= maponpathscomp (ezmap f g z (pr1 fs))
(hfiberpr1 g z)
(homotinvweqweq (ezweq f g z fs) x): maponpaths (hfiberpr1 g z)
(maponpaths (ezmap f g z (pr1 fs))
(homotinvweqweq (ezweq f g z fs) x)) =
maponpaths (hfiberpr1 g z ∘ ezmap f g z (pr1 fs))
(homotinvweqweq (ezweq f g z fs) x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e2: maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e3:= maponpaths
(λe : ezmap f g z (pr1 fs)
(invmap (ezweq f g z fs)
(ezmap f g z (pr1 fs) x)) =
ezmap f g z (pr1 fs) x,
maponpaths (hfiberpr1 g z) e)
(! homotweqinvweqweq (ezweq f g z fs) x): maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z (pr1 fs) x)) =
maponpaths (hfiberpr1 g z)
(maponpaths (ezmap f g z (pr1 fs))
(homotinvweqweq (ezweq f g z fs) x)) e4:= maponpathscomp (ezmap f g z (pr1 fs))
(hfiberpr1 g z)
(homotinvweqweq (ezweq f g z fs) x): maponpaths (hfiberpr1 g z)
(maponpaths (ezmap f g z (pr1 fs))
(homotinvweqweq (ezweq f g z fs) x)) =
maponpaths (hfiberpr1 g z ∘ ezmap f g z (pr1 fs))
(homotinvweqweq (ezweq f g z fs) x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e2: maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z x: X ff:= ezmap1 f g z fs (f x): g (f x) = z → hfiber f (f x) gg:= invezmap1 f g z (pr1 fs) (f x): hfiber f (f x) → g (f x) = z egf: ∏ e : g (f x) = z, gg (ff e) = e e1:= ! pathscomp0rid
(maponpaths f
(homotinvweqweq (ezweq f g z fs) x)): maponpaths f (homotinvweqweq (ezweq f g z fs) x) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x) e2: maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs)
(ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x)
maponpaths (hfiberpr1 g z)
(homotweqinvweq (ezweq f g z fs) (ezmap f g z fs x)) =
maponpaths f (homotinvweqweq (ezweq f g z fs) x) @
idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z fs: fibseqstr f g z y: Y ff:= ezmap1 f g z fs y: g y = z → hfiber f y gg:= invezmap1 f g z (pr1 fs) y: hfiber f y → g y = z egf: ∏ e : g y = z, gg (ff e) = e efg: ∏ xe : hfiber f y, ff (gg xe) = xe
isweq ff
apply (isweq_iso _ _ egf efg).Defined.Definitionezweq1 {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(fs : fibseqstr f g z) (y : Y) := make_weq _ (isweqezmap1 f g z fs y).Definitionfibseq1 {XYZ : UU} (f :X -> Y) (g : Y -> Z) (z : Z)
(fs : fibseqstr f g z) (y : Y) : fibseqstr (d1 f g z fs y) f y
:= make_fibseqstr _ _ _ _ (isweqezmap1 f g z fs y).(** *** Explicit description of the first map in the second derived sequence *)Definitiond2 {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(fs : fibseqstr f g z) (y : Y) (x : X) (e : (f x) = y) :
(g y) = z := pathscomp0 (maponpaths g (pathsinv0 e)) ((pr1 fs) x).Definitionezweq2 {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(fs : fibseqstr f g z) (y : Y) (x : X) :
f x = y ≃ hfiber (d1 f g z fs y) x
:= ezweq1 (d1 f g z fs y) f y (fibseq1 f g z fs y) x.Definitionfibseq2 {XYZ : UU} (f : X -> Y) (g : Y->Z) (z : Z)
(fs : fibseqstr f g z) (y : Y) (x : X) :
fibseqstr (d2 f g z fs y x) (d1 f g z fs y) x
:= make_fibseqstr _ _ _ _
(isweqezmap1 (d1 f g z fs y) f y (fibseq1 f g z fs y) x).(** *** Fibration sequences based on [ tpair P z ] and [ pr1 : total2 P -> Z ] ( the "pr1-case" ) *)(** Construction of the fibration sequence. *)Definitionezmappr1 {Z : UU} (P : Z -> UU) (z : Z) :
P z -> hfiber (@pr1 Z P) z
:= λp : P z, tpair _ (tpair _ z p) (idpath z).Definitioninvezmappr1 {Z : UU} (P : Z -> UU) (z : Z) :
hfiber (@pr1 Z P) z -> P z
:= λte, transportf P (pr2 te) (pr2 (pr1 te)).
Z: UU P: Z → UU z: Z
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z
∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z x: P z
invezmappr1 P z (ezmappr1 P z x) = x
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z x: P z
invezmappr1 P z ((z,, x),, idpath z) = x
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z x: P z
transportf P (pr2 ((z,, x),, idpath z))
(pr2 (pr1 ((z,, x),, idpath z))) = x
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z x: P z
transportf P (idpath z) x = x
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x
∏ x : hfiber pr1 z,
ezmappr1 P z (invezmappr1 P z x) = x
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x x: hfiber pr1 z
ezmappr1 P z (invezmappr1 P z x) = x
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x x: ∑ y, P y t0: pr1 x = z
ezmappr1 P z (invezmappr1 P z (x,, t0)) = x,, t0
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU x: ∑ y, P y egf: ∏ x0 : P (pr1 x), invezmappr1 P (pr1 x) (ezmappr1 P (pr1 x) x0) = x0
ezmappr1 P (pr1 x)
(invezmappr1 P (pr1 x) (x,, idpath (pr1 x))) =
x,, idpath (pr1 x)
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU x: ∑ y, P y egf: ∏ x0 : P (pr1 x), invezmappr1 P (pr1 x) (ezmappr1 P (pr1 x) x0) = x0
ezmappr1 P (pr1 x)
(invezmappr1 P (pr1 x) (x,, idpath (pr1 x))) =
x,, idpath (pr1 x)
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU x: ∑ y, P y egf: ∏ x0 : P (pr1 x), invezmappr1 P (pr1 x) (ezmappr1 P (pr1 x) x0) = x0
ezmappr1 P (pr1 x)
(invezmappr1 P (pr1 x) (x,, idpath (pr1 x))) =
x,, idpath (pr1 x)
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU pr1: Z pr2: P pr1 egf: ∏ x : P (Preamble.pr1 (pr1,, pr2)),
invezmappr1 P (Preamble.pr1 (pr1,, pr2))
(ezmappr1 P (Preamble.pr1 (pr1,, pr2)) x) = x
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU pr1: Z pr2: P pr1 egf: ∏ x : P (Preamble.pr1 (pr1,, pr2)),
invezmappr1 P (Preamble.pr1 (pr1,, pr2))
(ezmappr1 P (Preamble.pr1 (pr1,, pr2)) x) = x
ezmappr1 P pr1
(invezmappr1 P pr1 ((pr1,, pr2),, idpath pr1)) =
(pr1,, pr2),, idpath pr1
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU pr1: Z pr2: P pr1 egf: ∏ x : P (Preamble.pr1 (pr1,, pr2)),
invezmappr1 P (Preamble.pr1 (pr1,, pr2))
(ezmappr1 P (Preamble.pr1 (pr1,, pr2)) x) = x
ezmappr1 P pr1
(invezmappr1 P pr1 ((pr1,, pr2),, idpath pr1)) =
(pr1,, pr2),, idpath pr1
Z: UU P: Z → UU z: Z egf: ∏ x : P z, invezmappr1 P z (ezmappr1 P z x) = x efg: ∏ x : hfiber pr1 z, ezmappr1 P z (invezmappr1 P z x) = x
isweq (ezmappr1 P z)
Z: UU P: Z → UU pr1: Z pr2: P pr1 egf: ∏ x : P (Preamble.pr1 (pr1,, pr2)),
invezmappr1 P (Preamble.pr1 (pr1,, pr2))
(ezmappr1 P (Preamble.pr1 (pr1,, pr2)) x) = x
apply isweqezmappr1.Defined.Definitionfibseqpr1 {Z : UU} (P : Z -> UU) (z : Z) :
fibseqstr (λp : P z, tpair _ z p) (@pr1 Z P) z
:= make_fibseqstr _ _ _ _ (isfibseqpr1 P z).(** The main weak equivalence defined by the first derived of [ fibseqpr1 ]. *)Definitionezweq1pr1 {Z : UU} (P : Z -> UU) (z : Z) (zp : total2 P) :
pr1 zp = z ≃ hfiber (tpair P z) zp
:= ezweq1 _ _ z (fibseqpr1 P z) zp.(** *** Fibration sequences based on [ hfiberpr1 : hfiber g z -> Y ] and [ g : Y -> Z ] (the "g-case") *)
Y, Z: UU g: Y → Z z: Z
isfibseq (hfiberpr1 g z) g z
(λye : hfiber g z, pr2 ye)
Y, Z: UU g: Y → Z z: Z
isfibseq (hfiberpr1 g z) g z
(λye : hfiber g z, pr2 ye)
Y, Z: UU g: Y → Z z: Z
isfibseq (hfiberpr1 g z) g z
(λye : hfiber g z, pr2 ye)
Y, Z: UU g: Y → Z z: Z
idfun (hfiber g z) ~
ezmap (hfiberpr1 g z) g z (λye : hfiber g z, pr2 ye)
Y, Z: UU g: Y → Z z: Z
idfun (hfiber g z) ~
ezmap (hfiberpr1 g z) g z (λye : hfiber g z, pr2 ye)
Y, Z: UU g: Y → Z z: Z x: hfiber g z
idfun (hfiber g z) x =
ezmap (hfiberpr1 g z) g z (λye : hfiber g z, pr2 ye)
x
apply idpath.Defined.Definitionezweqg {YZ : UU} (g : Y -> Z) (z : Z) := make_weq _ (isfibseqg g z).Definitionfibseqg {YZ : UU} (g : Y -> Z) (z : Z)
: fibseqstr (hfiberpr1 g z) g z := make_fibseqstr _ _ _ _ (isfibseqg g z).(** The first derived of [ fibseqg ]. *)Definitiond1g {YZ : UU} (g : Y -> Z) (z : Z) (y : Y) :
(g y) = z -> hfiber g z := make_hfiber g y.(** note that [ d1g ] coincides with [ d1 _ _ _ (fibseqg g z) ] which makes the following two definitions possible. *)Definitionezweq1g {YZ : UU} (g : Y -> Z) (z : Z) (y : Y) :
g y = z ≃ hfiber (hfiberpr1 g z) y
:= make_weq _ (isweqezmap1 (hfiberpr1 g z) g z (fibseqg g z) y).Definitionfibseq1g {YZ : UU} (g : Y -> Z) (z : Z) (y : Y) :
fibseqstr (d1g g z y) (hfiberpr1 g z) y
:= make_fibseqstr _ _ _ _ (isweqezmap1 (hfiberpr1 g z) g z (fibseqg g z) y).(** The second derived of [ fibseqg ]. *)Definitiond2g {YZ : UU} (g : Y -> Z) {z : Z} (y : Y) (ye' : hfiber g z)
(e: (pr1 ye') = y) : (g y) = z
:= pathscomp0 (maponpaths g (pathsinv0 e)) (pr2 ye').(** note that [ d2g ] coincides with [ d2 _ _ _ (fibseqg g z) ] which makes the following two definitions possible. *)Definitionezweq2g
{YZ : UU} (g : Y -> Z) {z : Z} (y : Y) (ye' : hfiber g z) :
(pr1 ye') = y ≃ hfiber (make_hfiber g y) ye'
:= ezweq2 _ _ _ (fibseqg g z) _ _.Definitionfibseq2g {YZ : UU} (g : Y -> Z) {z : Z} (y : Y) (ye' : hfiber g z) :
fibseqstr (d2g g y ye') (make_hfiber g y) ye'
:= fibseq2 _ _ _ (fibseqg g z) _ _.(** The third derived of [ fibseqg ] and an explicit description of the corresponding first map. *)Definitiond3g {YZ : UU} (g : Y -> Z) {z : Z} (y : Y) (ye' : hfiber g z)
(e : (g y) = z) : (make_hfiber g y e) = ye' -> (pr1 ye') = y
:= d2 (d1g g z y) (hfiberpr1 g z) y (fibseq1g g z y) ye' e.
Y, Z: UU g: Y → Z z: Z y: Y ye': hfiber g z e: g y = z ee: make_hfiber g y e = ye'
d3g g y ye' e ee = maponpaths pr1 (! ee)
Y, Z: UU g: Y → Z z: Z y: Y ye': hfiber g z e: g y = z ee: make_hfiber g y e = ye'
d3g g y ye' e ee = maponpaths pr1 (! ee)
Y, Z: UU g: Y → Z z: Z y: Y ye': hfiber g z e: g y = z ee: make_hfiber g y e = ye'
d3g g y ye' e ee = maponpaths pr1 (! ee)
Y, Z: UU g: Y → Z z: Z y: Y ye': hfiber g z e: g y = z ee: make_hfiber g y e = ye'
d2 (d1g g z y) (hfiberpr1 g z) y (fibseq1g g z y) ye'
e ee = maponpaths pr1 (! ee)
Y, Z: UU g: Y → Z z: Z y: Y ye': hfiber g z e: g y = z ee: make_hfiber g y e = ye'
maponpaths (hfiberpr1 g z) (! ee) @
pr1 (fibseq1g g z y) e = maponpaths pr1 (! ee)
Y, Z: UU g: Y → Z z: Z y: Y ye': hfiber g z e: g y = z ee: make_hfiber g y e = ye'
maponpaths (hfiberpr1 g z) (! ee) @ idpath y =
maponpaths pr1 (! ee)
apply pathscomp0rid.Defined.Definitionezweq3g {YZ : UU} (g : Y -> Z) {z : Z} (y : Y) (ye' : hfiber g z)
(e : (g y) = z)
:= ezweq2 (d1g g z y) (hfiberpr1 g z) y (fibseq1g g z y) ye' e.Definitionfibseq3g {YZ : UU} (g : Y -> Z) {z : Z} (y : Y) (ye' : hfiber g z)
(e : (g y) = z)
:= fibseq2 (d1g g z y) (hfiberpr1 g z) y (fibseq1g g z y) ye' e.(** *** Fibration sequence of h-fibers defined by a composable pair of functions (the "hf-case") We construct a fibration sequence based on [ hfibersftogf f g z ye : hfiber f (pr1 ye) -> hfiber gf z) ] and [ hfibersgftog f g z : hfiber gf z -> hfiber g z) ].*)
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
hfiber (g ∘ f) z
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
hfiber (g ∘ f) z
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
hfiber (g ∘ f) z
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
(g ∘ f) (pr1 xe) = z
apply (pathscomp0 (maponpaths g (pr2 xe)) (pr2 ye)).Defined.
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
hfiber (hfibersgftog f g z) ye
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
hfiber (hfibersgftog f g z) ye
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
hfiber (hfibersgftog f g z) ye
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
hfibersgftog f g z (hfibersftogf f g z ye xe) = ye
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
hfibersgftog f g z (hfibersftogf f g z ye xe) = ye
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
pr2
(make_hfiber g (f (pr1 xe))
(maponpaths g (pr2 xe) @ pr2 ye)) =
maponpaths g (pr2 xe) @ pr2 ye
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xe: hfiber f (pr1 ye)
maponpaths g (pr2 xe) @ pr2 ye =
maponpaths g (pr2 xe) @ pr2 ye
apply idpath.Defined.
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xee': hfiber (hfibersgftog f g z) ye
hfiber f (pr1 ye)
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xee': hfiber (hfibersgftog f g z) ye
hfiber f (pr1 ye)
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xee': hfiber (hfibersgftog f g z) ye
hfiber f (pr1 ye)
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xee': hfiber (hfibersgftog f g z) ye
X, Y, Z: UU f: X → Y g: Y → Z y: Y x: X e: (g ∘ f) x = g y e': make_hfiber g (f (pr1 (x,, e))) (pr2 (x,, e)) =
y,, idpath (g y) e'':= maponpaths g (maponpaths (hfiberpr1 g (g y)) e'): g
(hfiberpr1 g (g y)
(make_hfiber g (f (pr1 (x,, e)))
(pr2 (x,, e)))) =
g (hfiberpr1 g (g y) (y,, idpath (g y)))
X, Y, Z: UU f: X → Y g: Y → Z y: Y x: X e: (g ∘ f) x = g y e': make_hfiber g (f x) e = y,, idpath (g y) e'':= maponpaths g (maponpaths (hfiberpr1 g (g y)) e'): g
(hfiberpr1 g (g y)
(make_hfiber g (f (pr1 (x,, e)))
(pr2 (x,, e)))) =
g (hfiberpr1 g (g y) (y,, idpath (g y)))
X, Y, Z: UU f: X → Y g: Y → Z y: Y x: X e: (g ∘ f) x = g y e': make_hfiber g (f x) e = y,, idpath (g y) e'':= maponpaths g (maponpaths (hfiberpr1 g (g y)) e'): g
(hfiberpr1 g (g y)
(make_hfiber g (f (pr1 (x,, e)))
(pr2 (x,, e)))) =
g (hfiberpr1 g (g y) (y,, idpath (g y)))
make_hfiber g
(f
(pr1 (make_hfiber (g ∘ f) x (e'' @ idpath (g y)))))
(pr2 (make_hfiber (g ∘ f) x (e'' @ idpath (g y)))) =
y,, idpath (g y)
X, Y, Z: UU f: X → Y g: Y → Z y: Y x: X e: (g ∘ f) x = g y e': make_hfiber g (f x) e = y,, idpath (g y) e'':= maponpaths g (maponpaths (hfiberpr1 g (g y)) e'): g
(hfiberpr1 g (g y)
(make_hfiber g (f (pr1 (x,, e)))
(pr2 (x,, e)))) =
g (hfiberpr1 g (g y) (y,, idpath (g y)))
apply (hfibertriangle2 _ (make_hfiber g (f x) (pathscomp0 e'' (idpath (g y))))
(make_hfiber g y (idpath _))
(maponpaths (hfiberpr1 _ _) e') (idpath _)).Defined.
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xee': hfiber (hfibersgftog f g z) ye
ffgg f g z ye xee' = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xee': hfiber (hfibersgftog f g z) ye
ffgg f g z ye xee' = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z xee': hfiber (hfibersgftog f g z) ye
ffgg f g z ye xee' = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z y: Y e: g y = z xee': hfiber (hfibersgftog f g z) (y,, e)
ffgg f g z (y,, e) xee' = xee'
X, Y, Z: UU f: X → Y g: Y → Z y: Y xee': hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
ffgg f g (g y) (y,, idpath (g y)) xee' = xee'
X, Y, Z: UU f: X → Y g: Y → Z y: Y xe: hfiber (g ∘ f) (g y) e': hfibersgftog f g (g y) xe = y,, idpath (g y)
ffgg f g (g y) (y,, idpath (g y)) (xe,, e') = xe,, e'
X, Y, Z: UU f: X → Y g: Y → Z y: Y xe: hfiber (g ∘ f) (g y)
ffgg f g (g y) (hfibersgftog f g (g y) xe)
(xe,, idpath (hfibersgftog f g (g y) xe)) =
xe,, idpath (hfibersgftog f g (g y) xe)
X, Y, Z: UU f: X → Y g: Y → Z y: Y x: X e: (g ∘ f) x = g y
ffgg f g (g y) (hfibersgftog f g (g y) (x,, e))
((x,, e),, idpath (hfibersgftog f g (g y) (x,, e))) =
(x,, e),, idpath (hfibersgftog f g (g y) (x,, e))
X, Y, Z: UU f: X → Y g: Y → Z y: Y x: X
ffgg f g ((g ∘ f) x)
(hfibersgftog f g ((g ∘ f) x)
(x,, idpath ((g ∘ f) x)))
((x,, idpath ((g ∘ f) x)),,
idpath
(hfibersgftog f g ((g ∘ f) x)
(x,, idpath ((g ∘ f) x)))) =
(x,, idpath ((g ∘ f) x)),,
idpath
(hfibersgftog f g ((g ∘ f) x)
(x,, idpath ((g ∘ f) x)))
X, Y, Z: UU f: X → Y g: Y → Z y: Y x: X
make_hfiber (g ∘ f) x (idpath (g (f x))),,
idpath (make_hfiber g (f x) (idpath (g (f x)))) =
(x,, idpath (g (f x))),,
idpath
(hfibersgftog f g (g (f x)) (x,, idpath (g (f x))))
apply idpath.Defined.
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z
isweq (ezmaphf f g z ye)
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z
isweq (ezmaphf f g z ye)
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z
isweq (ezmaphf f g z ye)
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye)
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye)
∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z y: Y e: g y = z ff:= ezmaphf f g z (y,, e): hfiber f (pr1 (y,, e))
→ hfiber (hfibersgftog f g z) (y,, e) gg:= invezmaphf f g z (y,, e): hfiber (hfibersgftog f g z) (y,, e)
→ hfiber f (pr1 (y,, e))
∏ xe : hfiber f (pr1 (y,, e)), gg (ff xe) = xe
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y)))
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) xe: hfiber f (pr1 (y,, idpath (g y)))
gg (ff xe) = xe
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) xe: hfiber f (pr1 (y,, idpath (g y)))
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) x: X ex: f x = pr1 (y,, idpath (g y))
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) x: X ex: f x = y
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z x: X ff:= ezmaphf f g (g (f x)) (f x,, idpath (g (f x))): hfiber f (pr1 (f x,, idpath (g (f x))))
→ hfiber (hfibersgftog f g (g (f x)))
(f x,, idpath (g (f x))) gg:= invezmaphf f g (g (f x)) (f x,, idpath (g (f x))): hfiber (hfibersgftog f g (g (f x)))
(f x,, idpath (g (f x)))
→ hfiber f (pr1 (f x,, idpath (g (f x))))
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z x: X ff:= ezmaphf f g (g (f x)) (f x,, idpath (g (f x))): hfiber f (pr1 (f x,, idpath (g (f x))))
→ hfiber (hfibersgftog f g (g (f x)))
(f x,, idpath (g (f x))) gg:= invezmaphf f g (g (f x)) (f x,, idpath (g (f x))): hfiber (hfibersgftog f g (g (f x)))
(f x,, idpath (g (f x)))
→ hfiber f (pr1 (f x,, idpath (g (f x))))
idpath (f x) = idpath (f x)
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe
∏ xee' : hfiber (hfibersgftog f g z) ye,
ff (gg xee') = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye, ff (gg xee') = xee'
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z y: Y e: g y = z ff:= ezmaphf f g z (y,, e): hfiber f (pr1 (y,, e))
→ hfiber (hfibersgftog f g z) (y,, e) gg:= invezmaphf f g z (y,, e): hfiber (hfibersgftog f g z) (y,, e)
→ hfiber f (pr1 (y,, e)) egf: ∏ xe : hfiber f (pr1 (y,, e)), gg (ff xe) = xe
∏ xee' : hfiber (hfibersgftog f g z) (y,, e),
ff (gg xee') = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye, ff (gg xee') = xee'
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))), gg (ff xe) = xe
∏
xee' : hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)), ff (gg xee') = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye, ff (gg xee') = xee'
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))), gg (ff xe) = xe xee': hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
ff (gg xee') = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye,
ff (gg xee') = xee'
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))),
gg (ff xe) = xe xee': hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
ff (gg xee') =
ffgg f g (g y) (make_hfiber g y (idpath (g y))) xee'
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))),
gg (ff xe) = xe xee': hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) hint: ff (gg xee') =
ffgg f g (g y) (make_hfiber g y (idpath (g y)))
xee'
ff (gg xee') = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye,
ff (gg xee') = xee'
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))),
gg (ff xe) = xe xe: hfiber (g ∘ f) (g y) e': hfibersgftog f g (g y) xe = y,, idpath (g y)
ff (gg (xe,, e')) =
ffgg f g (g y) (make_hfiber g y (idpath (g y)))
(xe,, e')
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))),
gg (ff xe) = xe xee': hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) hint: ff (gg xee') =
ffgg f g (g y) (make_hfiber g y (idpath (g y)))
xee'
ff (gg xee') = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye,
ff (gg xee') = xee'
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))),
gg (ff xe) = xe x: X e: (g ∘ f) x = g y e': hfibersgftog f g (g y) (x,, e) = y,, idpath (g y)
ff (gg ((x,, e),, e')) =
ffgg f g (g y) (make_hfiber g y (idpath (g y)))
((x,, e),, e')
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))),
gg (ff xe) = xe xee': hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) hint: ff (gg xee') =
ffgg f g (g y) (make_hfiber g y (idpath (g y)))
xee'
ff (gg xee') = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye,
ff (gg xee') = xee'
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z y: Y ff:= ezmaphf f g (g y) (y,, idpath (g y)): hfiber f (pr1 (y,, idpath (g y)))
→ hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) gg:= invezmaphf f g (g y) (y,, idpath (g y)): hfiber (hfibersgftog f g (g y))
(y,, idpath (g y))
→ hfiber f (pr1 (y,, idpath (g y))) egf: ∏ xe : hfiber f (pr1 (y,, idpath (g y))),
gg (ff xe) = xe xee': hfiber (hfibersgftog f g (g y))
(y,, idpath (g y)) hint: ff (gg xee') =
ffgg f g (g y) (make_hfiber g y (idpath (g y)))
xee'
ff (gg xee') = xee'
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye,
ff (gg xee') = xee'
isweq ff
X, Y, Z: UU f: X → Y g: Y → Z z: Z ye: hfiber g z ff:= ezmaphf f g z ye: hfiber f (pr1 ye)
→ hfiber (hfibersgftog f g z) ye gg:= invezmaphf f g z ye: hfiber (hfibersgftog f g z) ye
→ hfiber f (pr1 ye) egf: ∏ xe : hfiber f (pr1 ye), gg (ff xe) = xe efg: ∏ xee' : hfiber (hfibersgftog f g z) ye,
ff (gg xee') = xee'
isweq ff
apply (isweq_iso _ _ egf efg).Defined.Definitionezweqhf {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(ye : hfiber g z) :
hfiber f (pr1 ye) ≃ hfiber (hfibersgftog f g z) ye
:= make_weq _ (isweqezmaphf f g z ye).Definitionfibseqhf {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(ye : hfiber g z) :
fibseqstr (hfibersftogf f g z ye) (hfibersgftog f g z) ye
:= make_fibseqstr _ _ _ _ (isweqezmaphf f g z ye).Definitionisweqinvezmaphf {XYZ : UU} (f : X -> Y) (g : Y -> Z) (z : Z)
(ye : hfiber g z) :
isweq (invezmaphf f g z ye) := pr2 (invweq (ezweqhf f g z ye)).
X, Y, Z: UU w: X ≃ Y g: Y → Z z: Z
hfiber (g ∘ w) z ≃ hfiber g z
X, Y, Z: UU w: X ≃ Y g: Y → Z z: Z
hfiber (g ∘ w) z ≃ hfiber g z
X, Y, Z: UU w: X ≃ Y g: Y → Z z: Z
hfiber (g ∘ w) z ≃ hfiber g z
X, Y, Z: UU w: X ≃ Y g: Y → Z z: Z
isweq (hfibersgftog w g z)
X, Y, Z: UU w: X ≃ Y g: Y → Z z: Z ye: hfiber g z
iscontr (hfiber (hfibersgftog w g z) ye)
apply (iscontrweqf (ezweqhf w g z ye) ((pr2 w) (pr1 ye))).Defined.(** ** Functions between total spaces of families(** *** Function [ totalfun ] between total spaces from a family of functions between the fibersIncluding theorems saying that a fiber-wise morphism between total spaces is a weakequivalence if and only if all the morphisms between the fibers are weakequivalences. *)*)Definitiontotalfun {X : UU} (PQ : X -> UU) (f : ∏ x : X, P x -> Q x)
: (∑ y, P y) → (∑ y, Q y)
:= (λz: total2 P, tpair Q (pr1 z) (f (pr1 z) (pr2 z))).
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x
isweq (totalfun P Q f) → ∏ x : X, isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x
isweq (totalfun P Q f) → ∏ x : X, isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x
isweq hfx
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x
∏ y : hfiber piq x, iscontr (hfiber hfx y)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x y: hfiber piq x
iscontr (hfiber hfx y)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x y: hfiber piq x int:= invezmaphf totf piq x y: hfiber (hfibersgftog totf piq x) y
→ hfiber totf (pr1 y)
iscontr (hfiber hfx y)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x y: hfiber piq x int:= invezmaphf totf piq x y: hfiber (hfibersgftog totf piq x) y
→ hfiber totf (pr1 y)
isweq int
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x y: hfiber piq x int:= invezmaphf totf piq x y: hfiber (hfibersgftog totf piq x) y
→ hfiber totf (pr1 y) X1: isweq int
iscontr (hfiber hfx y)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x y: hfiber piq x int:= invezmaphf totf piq x y: hfiber (hfibersgftog totf piq x) y
→ hfiber totf (pr1 y) X1: isweq int
iscontr (hfiber hfx y)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x t: ∑ y : X, Q y e: piq t = x int:= invezmaphf totf piq x (t,, e): hfiber (hfibersgftog totf piq x) (t,, e)
→ hfiber totf (pr1 (t,, e)) X1: isweq int
iscontr (hfiber hfx (t,, e))
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x t: ∑ y : X, Q y e: piq t = x int:= invezmaphf totf piq x (t,, e): hfiber (hfibersgftog totf piq x) (t,, e)
→ hfiber totf (pr1 (t,, e)) X1: isweq int
iscontr (hfiber totf t)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x t: ∑ y : X, Q y e: piq t = x int:= invezmaphf totf piq x (t,, e): hfiber (hfibersgftog totf piq x) (t,, e)
→ hfiber totf (pr1 (t,, e)) X1: isweq int is1: iscontr (hfiber totf t)
iscontr (hfiber hfx (t,, e))
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x t: ∑ y : X, Q y e: piq t = x int:= invezmaphf totf piq x (t,, e): hfiber (hfibersgftog totf piq x) (t,, e)
→ hfiber totf (pr1 (t,, e)) X1: isweq int is1: iscontr (hfiber totf t)
iscontr (hfiber hfx (t,, e))
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x
isweq h
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x
∏ p : P x, h p = h' p
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x ee: ∏ p : P x, h p = h' p
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x p: P x
h p = h' p
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x ee: ∏ p : P x, h p = h' p
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x ee: ∏ p : P x, h p = h' p
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x ee: ∏ p : P x, h p = h' p
isweq h'
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x ee: ∏ p : P x, h p = h' p X2: isweq h'
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x ee: ∏ p : P x, h p = h' p X2: isweq h'
isweq (f x)
X: UU P, Q: X → UU f: ∏ x : X, P x → Q x X0: isweq (totalfun P Q f) x: X totp:= ∑ y, P y: Type totq:= ∑ y, Q y: Type totf:= totalfun P Q f: (∑ y : X, P y) → ∑ y : X, Q y pip:= λz : totp, pr1 z: totp → X piq:= λz : totq, pr1 z: totq → X hfx:= hfibersgftog totf piq x: hfiber (piq ∘ totf) x → hfiber piq x H: isweq hfx ip:= ezmappr1 P x: P x → hfiber pr1 x iq:= ezmappr1 Q x: Q x → hfiber pr1 x h:= λp : P x, hfx (ip p): P x → hfiber piq x is2: isweq h h':= λp : P x, iq (f x p): P x → hfiber pr1 x ee: ∏ p : P x, h p = h' p X2: isweq h'
isweq iq
apply (isweqezmappr1 Q x).Defined.Definitionweqtotaltofib {X : UU} (PQ : X -> UU) (f : ∏ x : X , P x -> Q x)
(is : isweq (totalfun _ _ f)) (x : X) : P x ≃ Q x
:= make_weq _ (isweqtotaltofib P Q f is x).
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
isweq (totalfun P Q (λx : X, f x))
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
isweq (totalfun P Q (λx : X, f x))
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y
isweq fpq
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X
isweq fpq
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X
isweq fpq
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X
∏ y : ∑ y : X, Q y, iscontr (hfiber fpq y)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x
isweq hfpqx
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint1: isweq hfpqx
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x ipx:= ezmappr1 P x: P x → hfiber pr1 x
isweq hfpqx
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint1: isweq hfpqx
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x ipx:= ezmappr1 P x: P x → hfiber pr1 x iqx:= ezmappr1 Q x: Q x → hfiber pr1 x
isweq hfpqx
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint1: isweq hfpqx
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x ipx:= ezmappr1 P x: P x → hfiber pr1 x iqx:= ezmappr1 Q x: Q x → hfiber pr1 x diag:= λp : P x, iqx (f x p): P x → hfiber pr1 x
isweq hfpqx
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint1: isweq hfpqx
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x ipx:= ezmappr1 P x: P x → hfiber pr1 x iqx:= ezmappr1 Q x: Q x → hfiber pr1 x diag:= λp : P x, iqx (f x p): P x → hfiber pr1 x
isweq diag
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x ipx:= ezmappr1 P x: P x → hfiber pr1 x iqx:= ezmappr1 Q x: Q x → hfiber pr1 x diag:= λp : P x, iqx (f x p): P x → hfiber pr1 x is2: isweq diag
isweq hfpqx
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint1: isweq hfpqx
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x ipx:= ezmappr1 P x: P x → hfiber pr1 x iqx:= ezmappr1 Q x: Q x → hfiber pr1 x diag:= λp : P x, iqx (f x p): P x → hfiber pr1 x is2: isweq diag
isweq hfpqx
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint1: isweq hfpqx
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint1: isweq hfpqx
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint1: ∏ y : hfiber pr1q x, iscontr (hfiber hfpqx y)
iscontr (hfiber hfpqx xqe)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe)
iscontr (hfiber fpq xq)
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x fpq:= totalfun P Q (λx : X, f x): (∑ y : X, P y) → ∑ y : X, Q y pr1p:= λz : ∑ y, P y, pr1 z: (∑ y, P y) → X pr1q:= λz : ∑ y, Q y, pr1 z: (∑ y, Q y) → X xq: ∑ y : X, Q y x:= pr1q xq: X xqe:= make_hfiber pr1q xq (idpath (pr1q xq)): hfiber pr1q (pr1q xq) hfpqx:= hfibersgftog fpq pr1q x: hfiber (pr1q ∘ fpq) x → hfiber pr1q x isint: iscontr (hfiber hfpqx xqe) intmap:= invezmaphf fpq pr1q x xqe: hfiber (hfibersgftog fpq pr1q x) xqe
→ hfiber fpq (pr1 xqe)
iscontr (hfiber fpq xq)
apply (iscontrweqf (make_weq intmap (isweqinvezmaphf fpq pr1q x xqe)) isint).Defined.
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
isweq (totalfun P Q (λx : X, f x))
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
isweq (totalfun P Q (λx : X, f x))
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
(∑ y : X, Q y) → ∑ y : X, P y
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
∏ x : ∑ y : X, P y,
?g (totalfun P Q (λx0 : X, f x0) x) = x
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
∏ y : ∑ y : X, Q y,
totalfun P Q (λx : X, f x) (?g y) = y
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
(∑ y : X, Q y) → ∑ y : X, P y
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
∏ x : X, Q x → P x
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x x: X
Q x → P x
apply (invmap (f x)).
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
∏ x : ∑ y : X, P y,
totalfun Q P (λx0 : X, invmap (f x0))
(totalfun P Q (λx0 : X, f x0) x) = x
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, P y
totalfun Q P (λx : X, invmap (f x))
(totalfun P Q (λx : X, f x) xp) = xp
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, P y
pr1
(totalfun Q P (λx : X, invmap (f x))
(totalfun P Q (λx : X, f x) xp)) = pr1 xp
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, P y
transportf P ?p
(pr2
(totalfun Q P (λx : X, invmap (f x))
(totalfun P Q (λx : X, f x) xp))) = pr2 xp
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, P y
pr1
(totalfun Q P (λx : X, invmap (f x))
(totalfun P Q (λx : X, f x) xp)) = pr1 xp
apply idpath.
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, P y
transportf P (idpath (pr1 xp))
(pr2
(totalfun Q P (λx : X, invmap (f x))
(totalfun P Q (λx : X, f x) xp))) = pr2 xp
apply homotinvweqweq.
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x
∏ y : ∑ y : X, Q y,
totalfun P Q (λx : X, f x)
(totalfun Q P (λx : X, invmap (f x)) y) = y
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, Q y
totalfun P Q (λx : X, f x)
(totalfun Q P (λx : X, invmap (f x)) xp) = xp
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, Q y
pr1
(totalfun P Q (λx : X, f x)
(totalfun Q P (λx : X, invmap (f x)) xp)) =
pr1 xp
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, Q y
transportf Q ?p
(pr2
(totalfun P Q (λx : X, f x)
(totalfun Q P (λx : X, invmap (f x)) xp))) =
pr2 xp
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, Q y
pr1
(totalfun P Q (λx : X, f x)
(totalfun Q P (λx : X, invmap (f x)) xp)) =
pr1 xp
apply idpath.
X: UU P, Q: X → UU f: ∏ x : X, P x ≃ Q x xp: ∑ y : X, Q y
transportf Q (idpath (pr1 xp))
(pr2
(totalfun P Q (λx : X, f x)
(totalfun Q P (λx : X, invmap (f x)) xp))) =
pr2 xp
apply homotweqinvweq.Defined.Definitionweqfibtototal {X : UU} (PQ : X -> UU) (f : ∏ x, P x ≃ Q x) :
(∑ x, P x) ≃ (∑ x, Q x) := make_weq _ (isweqfibtototal' P Q f).(** *** Function [ fpmap ] between the total spaces from a function between the basesGiven [ X Y ] in [ UU ], [ P:Y -> UU ] and [ f: X -> Y ] we get a function[ fpmap: total2 X (P f) -> total2 Y P ]. The main theorem of this sectionasserts that the homotopy fiber of fpmap over [ yp:total Y P ] is naturallyweakly equivalent to the homotopy fiber of [ f ] over [ pr1 yp ]. Inparticular, if [ f ] is a weak equivalence then so is [ fpmap ]. *)Definitionfpmap {XY : UU} (f : X -> Y) (P : Y -> UU) :
(∑ x, P (f x)) -> (∑ y, P y) := λz, tpair P (f (pr1 z)) (pr2 z).
X, Y: UU f: X → Y P: Y → UU
(∑ x : X, P (f x)) → ∑ u : ∑ y, P y, hfiber f (pr1 u)
X, Y: UU f: X → Y P: Y → UU
(∑ x : X, P (f x)) → ∑ u : ∑ y, P y, hfiber f (pr1 u)
X, Y: UU f: X → Y P: Y → UU X0: ∑ x : X, P (f x)
∑ u : ∑ y, P y, hfiber f (pr1 u)
X, Y: UU f: X → Y P: Y → UU X0: ∑ x : X, P (f x) u:= fpmap f P X0: ∑ y : Y, P y
∑ u : ∑ y, P y, hfiber f (pr1 u)
X, Y: UU f: X → Y P: Y → UU X0: ∑ x : X, P (f x) u:= fpmap f P X0: ∑ y : Y, P y
hfiber f (pr1 u)
X, Y: UU f: X → Y P: Y → UU X0: ∑ x : X, P (f x) u:= fpmap f P X0: ∑ y : Y, P y x:= pr1 X0: X
hfiber f (pr1 u)
X, Y: UU f: X → Y P: Y → UU X0: ∑ x : X, P (f x) u:= fpmap f P X0: ∑ y : Y, P y x:= pr1 X0: X
f x = pr1 u
X, Y: UU f: X → Y P: Y → UU X0: ∑ x : X, P (f x) u:= fpmap f P X0: ∑ y : Y, P y x:= pr1 X0: X
f x = f (pr1 X0)
apply idpath.Defined.
X: UU P: X → UU x: X
isweq (λp : P x, coconusfromtpair X (idpath x),, p)
X: UU P: X → UU x: X
isweq (λp : P x, coconusfromtpair X (idpath x),, p)
X: UU P: X → UU x: X
isweq (λp : P x, coconusfromtpair X (idpath x),, p)
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u)
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x
∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x z: ∑ u : coconusfromt X x, P (pr1 u)
f (g z) = z
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x t: coconusfromt X x x0: P (pr1 t)
f (g (t,, x0)) = t,, x0
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x t: X x1: x = t x0: P (pr1 (t,, x1))
f (g ((t,, x1),, x0)) = (t,, x1),, x0
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x t: X x1: x = t x0: P (pr1 (t,, x1))
f (g ((t,, x1),, x0)) = (t,, x1),, x0
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x x0: P (pr1 (x,, idpath x))
f (g ((x,, idpath x),, x0)) = (x,, idpath x),, x0
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x x0: P (pr1 (x,, idpath x))
f (g ((x,, idpath x),, x0)) = (x,, idpath x),, x0
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z
∏ p : P x, g (f p) = p
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z p: P x
g (f p) = p
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z egf: ∏ p : P x, g (f p) = p
isweq f
X: UU P: X → UU x: X f:= λp : P x, coconusfromtpair X (idpath x),, p: P x → ∑ u : coconusfromt X x, P (pr1 u) g:= λz : ∑ u : coconusfromt X x, P (pr1 u),
transportf P (! pr2 (pr1 z)) (pr2 z): (∑ u : coconusfromt X x, P (pr1 u)) → P x efg: ∏ z : ∑ u : coconusfromt X x, P (pr1 u), f (g z) = z egf: ∏ p : P x, g (f p) = p
isweq f
apply (isweq_iso f g egf efg).Defined.
X, Y: UU f: X → Y P: Y → UU
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u)
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u)
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u)
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u)
∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u)
∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u)
∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u)
∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) u: ∑ u : ∑ y, P y, hfiber f (pr1 u)
fromint (toint u) = u
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) t: ∑ y, P y x: hfiber f (pr1 t)
fromint (toint (t,, x)) = t,, x
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) t: ∑ y, P y pr1: X pr2: f pr1 = Preamble.pr1 t
fromint (toint (t,, pr1,, pr2)) = t,, pr1,, pr2
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) p0: Y p1: P p0 pr1: X pr2: f pr1 = Preamble.pr1 (p0,, p1)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) p0: Y p1: P p0 pr1: X pr2: f pr1 = Preamble.pr1 (p0,, p1)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) p0: Y p1: P p0 pr1: X pr2: f pr1 = Preamble.pr1 (p0,, p1)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) p0: Y p1: P p0 pr1: X pr2: f pr1 = Preamble.pr1 (p0,, p1)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) p0: Y p1: P p0 pr1: X pr2: f pr1 = Preamble.pr1 (p0,, p1)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u
∏ u : int, toint (fromint u) = u
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u u: int
toint (fromint u) = u
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u t: X x: ∑ u : coconusfromt Y (f t), P (pr1 u)
toint (fromint (t,, x)) = t,, x
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u t: X t0: coconusfromt Y (f t) x: P (pr1 t0)
toint (fromint (t,, t0,, x)) = t,, t0,, x
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) fromto: ∏
u : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
fromint (toint u) = u t: X pr1: Y pr2: f t = pr1 x: P (Preamble.pr1 (pr1,, pr2))
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) fromto: ∏
u : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
fromint (toint u) = u t: X pr1: Y pr2: f t = pr1 x: P pr1
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) fromto: ∏
u : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
fromint (toint u) = u t: X pr1: Y pr2: f t = pr1 x: P pr1
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) fromto: ∏
u : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
fromint (toint u) = u t: X pr1: Y pr2: f t = pr1 x: P pr1
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) fromto: ∏
u : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
fromint (toint u) = u t: X pr1: Y pr2: f t = pr1 x: P pr1
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
intpair (Preamble.pr1 (Preamble.pr2 z))
(coconusfromtpair Y
(Preamble.pr2 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr1 z)): (∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) fromint:= λz : int,
(Preamble.pr1 (Preamble.pr1 (Preamble.pr2 z)),,
Preamble.pr2 (Preamble.pr2 z)),,
make_hfiber f (Preamble.pr1 z)
(Preamble.pr2
(Preamble.pr1 (Preamble.pr2 z))): int
→ ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u) fromto: ∏
u : ∑ u : ∑ y, P y, hfiber f (Preamble.pr1 u),
fromint (toint u) = u t: X pr1: Y pr2: f t = pr1 x: P pr1
intpair t (coconusfromtpair Y pr2,, x) =
t,, (pr1,, pr2),, x
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u
isweq toint
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u is: isweq toint
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u tofrom: ∏ u : int, toint (fromint u) = u is: isweq toint
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) fromto: ∏ u : ∑ u : ∑ y, P y, hfiber f (pr1 u),
fromint (toint u) = u is: isweq toint
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) fromint:= λz : int,
(pr1 (pr1 (pr2 z)),, pr2 (pr2 z)),,
make_hfiber f (pr1 z) (pr2 (pr1 (pr2 z))): int → ∑ u : ∑ y, P y, hfiber f (pr1 u) is: isweq toint
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u)
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u)
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u)
∏ x : X,
isweq
(λp : P (f x),
coconusfromtpair Y (idpath (f x)),, p)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u) l1: ∏ x : X,
isweq
(λp : P (f x),
coconusfromtpair Y (idpath (f x)),, p)
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u) x: X
isweq
(λp : P (f x),
coconusfromtpair Y (idpath (f x)),, p)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u) l1: ∏ x : X,
isweq
(λp : P (f x),
coconusfromtpair Y (idpath (f x)),, p)
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u) l1: ∏ x : X,
isweq
(λp : P (f x),
coconusfromtpair Y (idpath (f x)),, p)
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u) l1: ∏ x : X,
isweq
(λp : P (f x),
coconusfromtpair Y (idpath (f x)),, p)
isweq h
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u) l1: ∏ x : X,
isweq
(λp : P (f x),
coconusfromtpair Y (idpath (f x)),, p) X0: isweq h
isweq (hffpmap2 f P)
X, Y: UU f: X → Y P: Y → UU int:= ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u): Type intpair:= tpair
(λx : X,
∑ u : coconusfromt Y (f x), P (pr1 u)): ∏ pr1 : X,
(λx : X,
∑ u : coconusfromt Y (f x),
P (Preamble.pr1 u)) pr1
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (Preamble.pr1 u) toint:= λz : ∑ u : ∑ y, P y, hfiber f (pr1 u),
intpair (pr1 (pr2 z))
(coconusfromtpair Y (pr2 (pr2 z)),,
pr2 (pr1 z)): (∑ u : ∑ y, P y, hfiber f (pr1 u))
→ ∑ (x : X) (u : coconusfromt Y (f x)),
P (pr1 u) is: isweq toint h:= λu : ∑ x : X, P (f x), toint (hffpmap2 f P u): (∑ x : X, P (f x))
→ ∑ (x : X) (u : coconusfromt Y (f x)), P (pr1 u) l1: ∏ x : X,
isweq
(λp : P (f x),
coconusfromtpair Y (idpath (f x)),, p) X0: isweq h
isweq (hffpmap2 f P)
apply (twooutof3a (hffpmap2 f P) toint X0 is).Defined.(** *** Homotopy fibers of [ fpmap ] *)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y
hfiber (fpmap f P) yp → hfiber f (pr1 yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y
hfiber (fpmap f P) yp → hfiber f (pr1 yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y X0: hfiber (fpmap f P) yp
hfiber f (pr1 yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y X0: hfiber (fpmap f P) yp int1:= hfibersgftog (hffpmap2 f P) (λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp
hfiber f (pr1 yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y X0: hfiber (fpmap f P) yp int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp phi:= invezmappr1 (λu : ∑ y, P y, hfiber f (pr1 u)) yp: hfiber pr1 yp
→ (λu : ∑ y, P y, hfiber f (pr1 u)) yp
hfiber f (pr1 yp)
apply (phi (int1 X0)).Defined.
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P) (λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp
isweq int1
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp is1: isweq int1
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp
isweq int1
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp is1: isweq int1
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp is1: isweq int1
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp is1: isweq int1 phi:= invezmappr1 (λu : ∑ y, P y, hfiber f (pr1 u)) yp: hfiber pr1 yp
→ (λu : ∑ y, P y, hfiber f (pr1 u)) yp
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp is1: isweq int1 phi:= invezmappr1 (λu : ∑ y, P y, hfiber f (pr1 u)) yp: hfiber pr1 yp
→ (λu : ∑ y, P y, hfiber f (pr1 u)) yp
isweq phi
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp is1: isweq int1 phi:= invezmappr1 (λu : ∑ y, P y, hfiber f (pr1 u)) yp: hfiber pr1 yp
→ (λu : ∑ y, P y, hfiber f (pr1 u)) yp is2: isweq phi
isweq (hfiberfpmap f P yp)
X, Y: UU f: X → Y P: Y → UU yp: ∑ y, P y int1:= hfibersgftog (hffpmap2 f P)
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u), pr1 u)
yp: hfiber
((λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) ∘ hffpmap2 f P) yp
→ hfiber
(λu : ∑ u : ∑ y, P y, hfiber f (pr1 u),
pr1 u) yp is1: isweq int1 phi:= invezmappr1 (λu : ∑ y, P y, hfiber f (pr1 u)) yp: hfiber pr1 yp
→ (λu : ∑ y, P y, hfiber f (pr1 u)) yp is2: isweq phi
isweq (hfiberfpmap f P yp)
apply (twooutof3c int1 phi is1 is2).Defined.(** *** The [ fpmap ] from a weak equivalence is a weak equivalence *)
X, Y: UU w: X ≃ Y P: Y → UU
isweq (fpmap w P)
X, Y: UU w: X ≃ Y P: Y → UU
isweq (fpmap w P)
X, Y: UU w: X ≃ Y P: Y → UU
isweq (fpmap w P)
X, Y: UU w: X ≃ Y P: Y → UU
∏ y : ∑ y : Y, P y, iscontr (hfiber (fpmap w P) y)
X, Y: UU w: X ≃ Y P: Y → UU y: ∑ y : Y, P y
iscontr (hfiber (fpmap w P) y)
X, Y: UU w: X ≃ Y P: Y → UU y: ∑ y : Y, P y h:= hfiberfpmap w P y: hfiber (fpmap w P) y → hfiber w (pr1 y)
iscontr (hfiber (fpmap w P) y)
X, Y: UU w: X ≃ Y P: Y → UU y: ∑ y : Y, P y h:= hfiberfpmap w P y: hfiber (fpmap w P) y → hfiber w (pr1 y)
isweq h
X, Y: UU w: X ≃ Y P: Y → UU y: ∑ y : Y, P y h:= hfiberfpmap w P y: hfiber (fpmap w P) y → hfiber w (pr1 y) X1: isweq h
iscontr (hfiber (fpmap w P) y)
X, Y: UU w: X ≃ Y P: Y → UU y: ∑ y : Y, P y h:= hfiberfpmap w P y: hfiber (fpmap w P) y → hfiber w (pr1 y) X1: isweq h
iscontr (hfiber (fpmap w P) y)
X, Y: UU w: X ≃ Y P: Y → UU y: ∑ y : Y, P y h:= hfiberfpmap w P y: hfiber (fpmap w P) y → hfiber w (pr1 y) X1: isweq h
iscontr (hfiber w (pr1 y))
X, Y: UU w: X ≃ Y P: Y → UU y: ∑ y : Y, P y h:= hfiberfpmap w P y: hfiber (fpmap w P) y → hfiber w (pr1 y) X1: isweq h is: iscontr (hfiber w (pr1 y))
iscontr (hfiber (fpmap w P) y)
X, Y: UU w: X ≃ Y P: Y → UU y: ∑ y : Y, P y h:= hfiberfpmap w P y: hfiber (fpmap w P) y → hfiber w (pr1 y) X1: isweq h is: iscontr (hfiber w (pr1 y))
iscontr (hfiber (fpmap w P) y)
apply (iscontrweqb (make_weq h X1) is).Defined.
X, Y: UU w: X ≃ Y P: Y → UU
(∑ x : X, P (w x)) → ∑ y : Y, P y
X, Y: UU w: X ≃ Y P: Y → UU
(∑ x : X, P (w x)) → ∑ y : Y, P y
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
∑ y : Y, P y
exact (w (pr1 xp),,pr2 xp).Defined.
X, Y: UU w: X ≃ Y P: Y → UU
(∑ y : Y, P y) → ∑ x : X, P (w x)
X, Y: UU w: X ≃ Y P: Y → UU
(∑ y : Y, P y) → ∑ x : X, P (w x)
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
∑ x : X, P (w x)
exact (invmap w (pr1 yp),,
transportf P (! homotweqinvweq w (pr1 yp)) (pr2 yp)).Defined.
X, Y: UU w: X ≃ Y P: Y → UU
(∑ x : X, P (w x)) ≃ (∑ y : Y, P y)
X, Y: UU w: X ≃ Y P: Y → UU
(∑ x : X, P (w x)) ≃ (∑ y : Y, P y)
X, Y: UU w: X ≃ Y P: Y → UU
(∑ x : X, P (w x)) ≃ (∑ y : Y, P y)
X, Y: UU w: X ≃ Y P: Y → UU
isweq (weqfp_map w P)
X, Y: UU w: X ≃ Y P: Y → UU
∏ x : ∑ x : X, P (w x),
weqfp_invmap w P (weqfp_map w P x) = x
X, Y: UU w: X ≃ Y P: Y → UU
∏ y : ∑ y : Y, P y,
weqfp_map w P (weqfp_invmap w P y) = y
X, Y: UU w: X ≃ Y P: Y → UU
∏ x : ∑ x : X, P (w x),
weqfp_invmap w P (weqfp_map w P x) = x
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
weqfp_invmap w P (weqfp_map w P xp) = xp
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
pr1 (weqfp_invmap w P (weqfp_map w P xp)) = pr1 xp
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
transportf (λx : X, P (w x)) ?p
(pr2 (weqfp_invmap w P (weqfp_map w P xp))) = pr2 xp
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
pr1 (weqfp_invmap w P (weqfp_map w P xp)) = pr1 xp
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
invmap w (w (pr1 xp)) = pr1 xp
apply homotinvweqweq.
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
transportf (λx : X, P (w x))
(homotinvweqweq w (pr1 xp)
:
pr1 (weqfp_invmap w P (weqfp_map w P xp)) = pr1 xp)
(pr2 (weqfp_invmap w P (weqfp_map w P xp))) = pr2 xp
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
transportf (λx : X, P (w x))
(homotinvweqweq w (pr1 xp))
(transportf P (! homotweqinvweq w (w (pr1 xp)))
(pr2 xp)) = pr2 xp
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
transportf (λx : X, P (w x))
(homotinvweqweq w (pr1 xp))
(transportf (P ∘ w) (! homotinvweqweq w (pr1 xp))
(pr2 xp)) = pr2 xp
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
transportf (λx : X, P (w x))
(! homotinvweqweq w (pr1 xp) @
homotinvweqweq w (pr1 xp)) (pr2 xp) = pr2 xp
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
transportf (λx : X, P (w x)) (idpath (pr1 xp))
(pr2 xp) = pr2 xp
apply idpath.
X, Y: UU w: X ≃ Y P: Y → UU
∏ y : ∑ y : Y, P y,
weqfp_map w P (weqfp_invmap w P y) = y
X, Y: UU w: X ≃ Y P: Y → UU
∏ y : ∑ y : Y, P y,
weqfp_map w P (weqfp_invmap w P y) = y
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
weqfp_map w P (weqfp_invmap w P yp) = yp
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
pr1 (weqfp_map w P (weqfp_invmap w P yp)) = pr1 yp
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
transportf P ?p
(pr2 (weqfp_map w P (weqfp_invmap w P yp))) = pr2 yp
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
pr1 (weqfp_map w P (weqfp_invmap w P yp)) = pr1 yp
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
w (invmap w (pr1 yp)) = pr1 yp
apply homotweqinvweq.
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
transportf P
(homotweqinvweq w (pr1 yp)
:
pr1 (weqfp_map w P (weqfp_invmap w P yp)) = pr1 yp)
(pr2 (weqfp_map w P (weqfp_invmap w P yp))) = pr2 yp
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
transportf P (homotweqinvweq w (pr1 yp))
(transportf P (! homotweqinvweq w (pr1 yp)) (pr2 yp)) =
pr2 yp
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
transportf P
(! homotweqinvweq w (pr1 yp) @
homotweqinvweq w (pr1 yp)) (pr2 yp) = pr2 yp
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
transportf P (idpath (pr1 yp)) (pr2 yp) = pr2 yp
apply idpath.}Defined.
X, Y: UU w: X ≃ Y P: Y → UU
weqfp w P ~ weqfp_map w P
X, Y: UU w: X ≃ Y P: Y → UU
weqfp w P ~ weqfp_map w P
X, Y: UU w: X ≃ Y P: Y → UU
weqfp w P ~ weqfp_map w P
X, Y: UU w: X ≃ Y P: Y → UU xp: ∑ x : X, P (w x)
weqfp w P xp = weqfp_map w P xp
apply idpath.Defined.
X, Y: UU w: X ≃ Y P: Y → UU
invmap (weqfp w P) ~ weqfp_invmap w P
X, Y: UU w: X ≃ Y P: Y → UU
invmap (weqfp w P) ~ weqfp_invmap w P
X, Y: UU w: X ≃ Y P: Y → UU
invmap (weqfp w P) ~ weqfp_invmap w P
X, Y: UU w: X ≃ Y P: Y → UU yp: ∑ y : Y, P y
invmap (weqfp w P) yp = weqfp_invmap w P yp
apply idpath.Defined.
W, X, Y: UU P: W → UU f: X ⨿ Y ≃ W
(∑ w : W, P w)
≃ (∑ x : X, P (f (inl x))) ⨿ (∑ y : Y, P (f (inr y)))
W, X, Y: UU P: W → UU f: X ⨿ Y ≃ W
(∑ w : W, P w)
≃ (∑ x : X, P (f (inl x))) ⨿ (∑ y : Y, P (f (inr y)))
W, X, Y: UU P: W → UU f: X ⨿ Y ≃ W
(∑ w : W, P w)
≃ (∑ x : X, P (f (inl x))) ⨿ (∑ y : Y, P (f (inr y)))
exact (weqcomp (invweq (weqfp f _)) (weqtotal2overcoprod (P ∘ f))).Defined.(** *** Total spaces of families over a contractible base *)
P: unit → UU tp: ∑ y, P y
P tt
P: unit → UU tp: ∑ y, P y
P tt
P: unit → UU tp: ∑ y, P y
P tt
P: unit → UU t: unit p: P t
P tt
P: unit → UU p: P tt
P tt
apply p.Defined.Definitiontototal2overunit (P : unit -> UU) (p : P tt) : total2 P
:= tpair P tt p.
P: unit → UU
(∑ u : unit, P u) ≃ P tt
P: unit → UU
(∑ u : unit, P u) ≃ P tt
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt
(∑ u : unit, P u) ≃ P tt
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y
(∑ u : unit, P u) ≃ P tt
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y
isweq f
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y
∏ a : ∑ y, P y, g (f a) = a
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
isweq f
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y a: ∑ y, P y
g (f a) = a
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
isweq f
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y t: unit p: P t
g (f (t,, p)) = t,, p
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
isweq f
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y p: P tt
g (f (tt,, p)) = tt,, p
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
isweq f
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
isweq f
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a
∏ a : P tt, f (g a) = a
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a efg: ∏ a : P tt, f (g a) = a
isweq f
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a a: P tt
f (g a) = a
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a efg: ∏ a : P tt, f (g a) = a
isweq f
P: unit → UU f:= fromtotal2overunit P: (∑ y, P y) → P tt g:= tototal2overunit P: P tt → ∑ y, P y egf: ∏ a : ∑ y, P y, g (f a) = a efg: ∏ a : P tt, f (g a) = a
isweq f
apply (isweq_iso _ _ egf efg).Defined.(** *** The function on the total spaces from functions on the bases and on the fibers *)Definitionbandfmap {XY : UU} (f : X -> Y) (P : X -> UU) (Q : Y -> UU)
(fm : ∏ x : X, P x -> (Q (f x))) :
(∑ x, P x) -> (∑ x, Q x) := λxp, f (pr1 xp) ,, fm (pr1 xp) (pr2 xp).
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x)
isweq (bandfmap w P Q (λx : X, fw x))
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x)
isweq (bandfmap w P Q (λx : X, fw x))
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x)
isweq (bandfmap w P Q (λx : X, fw x))
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y
isweq (bandfmap w P Q (λx : X, fw x))
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y is1:= isweqfibtototal P (λx : X, Q (w x)) fw: isweq (totalfun P (λx : X, Q (w x)) (λx : X, fw x))
isweq (bandfmap w P Q (λx : X, fw x))
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y is1:= isweqfibtototal P (λx : X, Q (w x)) fw: isweq (totalfun P (λx : X, Q (w x)) (λx : X, fw x)) f2:= fpmap w Q: (∑ x : X, Q (w x)) → ∑ y : Y, Q y
isweq (bandfmap w P Q (λx : X, fw x))
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y is1:= isweqfibtototal P (λx : X, Q (w x)) fw: isweq (totalfun P (λx : X, Q (w x)) (λx : X, fw x)) f2:= fpmap w Q: (∑ x : X, Q (w x)) → ∑ y : Y, Q y is2:= isweqfpmap w Q: isweq (fpmap w Q)
isweq (bandfmap w P Q (λx : X, fw x))
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y is1:= isweqfibtototal P (λx : X, Q (w x)) fw: isweq (totalfun P (λx : X, Q (w x)) (λx : X, fw x)) f2:= fpmap w Q: (∑ x : X, Q (w x)) → ∑ y : Y, Q y is2:= isweqfpmap w Q: isweq (fpmap w Q)
∏ xp : ∑ y, P y,
f2 (f1 xp) = bandfmap w P Q (λx : X, fw x) xp
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y is1:= isweqfibtototal P (λx : X, Q (w x)) fw: isweq (totalfun P (λx : X, Q (w x)) (λx : X, fw x)) f2:= fpmap w Q: (∑ x : X, Q (w x)) → ∑ y : Y, Q y is2:= isweqfpmap w Q: isweq (fpmap w Q) h: ∏ xp : ∑ y, P y,
f2 (f1 xp) = bandfmap w P Q (λx : X, fw x) xp
isweq (bandfmap w P Q (λx : X, fw x))
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y is1:= isweqfibtototal P (λx : X, Q (w x)) fw: isweq (totalfun P (λx : X, Q (w x)) (λx : X, fw x)) f2:= fpmap w Q: (∑ x : X, Q (w x)) → ∑ y : Y, Q y is2:= isweqfpmap w Q: isweq (fpmap w Q)
∏ xp : ∑ y, P y,
f2 (f1 xp) = bandfmap w P Q (λx : X, fw x) xp
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y is1:= isweqfibtototal P (λx : X, Q (w x)) fw: isweq (totalfun P (λx : X, Q (w x)) (λx : X, fw x)) f2:= fpmap w Q: (∑ x : X, Q (w x)) → ∑ y : Y, Q y is2:= isweqfpmap w Q: isweq (fpmap w Q) xp: ∑ y, P y
f2 (f1 xp) = bandfmap w P Q (λx : X, fw x) xp
apply idpath.
X, Y: UU w: X ≃ Y P: X → UU Q: Y → UU fw: ∏ x : X, P x ≃ Q (w x) f1:= totalfun P (λx : X, Q (w x)) (λx : X, fw x): (∑ y : X, P y) → ∑ y : X, (λx : X, Q (w x)) y is1:= isweqfibtototal P (λx : X, Q (w x)) fw: isweq (totalfun P (λx : X, Q (w x)) (λx : X, fw x)) f2:= fpmap w Q: (∑ x : X, Q (w x)) → ∑ y : Y, Q y is2:= isweqfpmap w Q: isweq (fpmap w Q) h: ∏ xp : ∑ y, P y,
f2 (f1 xp) = bandfmap w P Q (λx : X, fw x) xp
isweq (bandfmap w P Q (λx : X, fw x))
apply (isweqhomot _ _ h (twooutof3c f1 f2 is1 is2)).Defined.Definitionweqbandf {XY : UU} (w : X ≃ Y) (P : X -> UU) (Q : Y -> UU)
(fw : ∏ x : X, P x ≃ Q (w x))
:= make_weq _ (isweqbandfmap w P Q fw).(** ** Homotopy fiber squares *)(** *** Homotopy commutative squares *)Definitioncommsqstr {XX'YZ : UU} (g' : Z -> X') (f' : X' -> Y) (g : Z -> X)
(f : X -> Y) := ∏ (z : Z), (f' (g' z)) = (f (g z)).
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x: X ze: hfiber g x
hfiber f' (f x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x: X ze: hfiber g x
hfiber f' (f x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x: X ze: hfiber g x
hfiber f' (f x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x: X z: Z e: g z = x
hfiber f' (f x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x: X z: Z e: g z = x
f' (g' z) = f x
apply (pathscomp0 (h z) (maponpaths f e)).Defined.
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x': X' ze: hfiber g' x'
hfiber f (f' x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x': X' ze: hfiber g' x'
hfiber f (f' x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x': X' ze: hfiber g' x'
hfiber f (f' x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x': X' z: Z e: g' z = x'
hfiber f (f' x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f x': X' z: Z e: g' z = x'
f (g z) = f' x'
apply (pathscomp0 (pathsinv0 (h z)) (maponpaths f' e)).Defined.Definitiontransposcommsqstr {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X') :
commsqstr g' f' g f -> commsqstr g f g' f'
:= λh : _, λz : Z, (pathsinv0 (h z)).(** *** Short complexes and homotopy commutative squares *)
X, Y, Z: UU f: X → Y g: Y → Z z: Z h: complxstr f g z
commsqstr f g (λ_ : X, tt) (λ_ : unit, z)
X, Y, Z: UU f: X → Y g: Y → Z z: Z h: complxstr f g z
commsqstr f g (λ_ : X, tt) (λ_ : unit, z)
X, Y, Z: UU f: X → Y g: Y → Z z: Z h: complxstr f g z
commsqstr f g (λ_ : X, tt) (λ_ : unit, z)
assumption.Defined.
X, Y, Z: UU f: X → Y g: Y → Z z: Z h: commsqstr f g (λ_ : X, tt) (λ_ : unit, z)
complxstr f g z
X, Y, Z: UU f: X → Y g: Y → Z z: Z h: commsqstr f g (λ_ : X, tt) (λ_ : unit, z)
complxstr f g z
X, Y, Z: UU f: X → Y g: Y → Z z: Z h: commsqstr f g (λ_ : X, tt) (λ_ : unit, z)
complxstr f g z
assumption.Defined.(** *** Homotopy fiber products *)Definitionhfp {XX'Y : UU} (f : X -> Y) (f' : X' -> Y)
:= total2 (λxx' : X × X', (f' (pr2 xx')) = (f (pr1 xx'))).Definitionhfpg {XX'Y : UU} (f : X -> Y) (f' : X' -> Y) :
hfp f f' -> X := λxx'e, (pr1 (pr1 xx'e)).Definitionhfpg' {XX'Y : UU} (f : X -> Y) (f' : X' -> Y) :
hfp f f' -> X' := λxx'e, (pr2 (pr1 xx'e)).DefinitioncommsqZtohfp {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X') (h : commsqstr g' f' g f) :
Z -> hfp f f' := λz : _, tpair _ (make_dirprod (g z) (g' z)) (h z).DefinitioncommsqZtohfphomot {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X') (h : commsqstr g' f' g f) :
∏ z : Z, (hfpg _ _ (commsqZtohfp _ _ _ _ h z)) = (g z)
:= λz : _, idpath _.DefinitioncommsqZtohfphomot' {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X') (h : commsqstr g' f' g f) :
∏ z : Z, (hfpg' _ _ (commsqZtohfp _ _ _ _ h z)) = (g' z)
:= λz : _, idpath _.DefinitionhfpoverX {XX'Y : UU} (f : X -> Y) (f' : X' -> Y)
:= total2 (λx : X, hfiber f' (f x)).DefinitionhfpoverX' {XX'Y : UU} (f : X -> Y) (f' : X' -> Y)
:= total2 (λx' : X', hfiber f (f' x')).
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfpoverX f f'
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfpoverX f f'
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfpoverX f f'
exact (weqtotal2asstor _ (λxx', f' (pr2 xx') = f (pr1 xx'))).Defined.
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfpoverX' f f'
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfpoverX' f f'
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfpoverX' f f'
X, X', Y: UU f: X → Y f': X' → Y w1:= weqfp (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')): (∑ x : X × X',
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(weqdirprodcomm X X' x))
≃ (∑ y : X' × X,
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')) y)
hfp f f' ≃ hfpoverX' f f'
X, X', Y: UU f: X → Y f': X' → Y w1:= weqfp (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')): (∑ x : X × X', f' (pr2 x) = f (pr1 x))
≃ (∑ y : X' × X, f' (pr1 y) = f (pr2 y))
hfp f f' ≃ hfpoverX' f f'
X, X', Y: UU f: X → Y f': X' → Y w1:= weqfp (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')): (∑ x : X × X', f' (pr2 x) = f (pr1 x))
≃ (∑ y : X' × X, f' (pr1 y) = f (pr2 y)) w2:= weqfibtototal
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx'x : X' × X,
weqpathsinv0 (f' (pr1 x'x)) (f (pr2 x'x))): (∑ x : X' × X,
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x)) x)
≃ (∑ x : X' × X,
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x)) x)
hfp f f' ≃ hfpoverX' f f'
X, X', Y: UU f: X → Y f': X' → Y w1:= weqfp (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')): (∑ x : X × X', f' (pr2 x) = f (pr1 x))
≃ (∑ y : X' × X, f' (pr1 y) = f (pr2 y)) w2:= weqfibtototal
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx'x : X' × X,
weqpathsinv0 (f' (pr1 x'x)) (f (pr2 x'x))): (∑ x : X' × X,
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x)) x)
≃ (∑ x : X' × X,
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x)) x) w3:= weqtotal2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x)): (∑ x'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
≃ (∑ (x : X') (p : (λ_ : X', X) x),
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(x,, p))
hfp f f' ≃ hfpoverX' f f'
X, X', Y: UU f: X → Y f': X' → Y w1:= weqfp (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')): (∑ x : X × X', f' (pr2 x) = f (pr1 x))
≃ (∑ y : X' × X, f' (pr1 y) = f (pr2 y)) w2:= weqfibtototal
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx'x : X' × X,
weqpathsinv0 (f' (pr1 x'x)) (f (pr2 x'x))): (∑ x : X' × X,
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x)) x)
≃ (∑ x : X' × X,
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x)) x) w3:= weqtotal2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x)): (∑ x'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
≃ (∑ (x : X') (p : X), f p = f' x)
hfp f f' ≃ hfpoverX' f f'
apply (weqcomp (weqcomp w1 w2) w3).Defined.
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfp f' f
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfp f' f
X, X', Y: UU f: X → Y f': X' → Y
hfp f f' ≃ hfp f' f
X, X', Y: UU f: X → Y f': X' → Y w1:= weqfp (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')): (∑ x : X × X',
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(weqdirprodcomm X X' x))
≃ (∑ y : X' × X,
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')) y)
hfp f f' ≃ hfp f' f
X, X', Y: UU f: X → Y f': X' → Y w1:= weqfp (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')): (∑ x : X × X', f' (pr2 x) = f (pr1 x))
≃ (∑ y : X' × X, f' (pr1 y) = f (pr2 y))
hfp f f' ≃ hfp f' f
X, X', Y: UU f: X → Y f': X' → Y w1:= weqfp (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx')): (∑ x : X × X', f' (pr2 x) = f (pr1 x))
≃ (∑ y : X' × X, f' (pr1 y) = f (pr2 y)) w2:= weqfibtototal
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx'x : X' × X,
weqpathsinv0 (f' (pr1 x'x)) (f (pr2 x'x))): (∑ x : X' × X,
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x)) x)
≃ (∑ x : X' × X,
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x)) x)
hfp f f' ≃ hfp f' f
apply (weqcomp w1 w2).Defined.Definitioncommhfp {XX'Y : UU} (f : X -> Y) (f' : X' -> Y) :
commsqstr (hfpg' f f') f' (hfpg f f') f
:= λxx'e : hfp f f', pr2 xx'e.(** *** Homotopy fiber products and homotopy fibers *)Definitionhfibertohfp {XY : UU} (f : X -> Y) (y : Y) (xe : hfiber f y) :
hfp (λt : unit, y) f := tpair (λtx : unit × X, (f (pr2 tx)) = y)
(make_dirprod tt (pr1 xe)) (pr2 xe).Definitionhfptohfiber {XY : UU} (f : X -> Y) (y : Y)
(hf : hfp (λt : unit, y) f) : hfiber f y
:= make_hfiber f (pr2 (pr1 hf)) (pr2 hf).
X, Y: UU f: X → Y y: Y
hfiber f y ≃ hfp (λ_ : unit, y) f
X, Y: UU f: X → Y y: Y
hfiber f y ≃ hfp (λ_ : unit, y) f
X, Y: UU f: X → Y y: Y
hfiber f y ≃ hfp (λ_ : unit, y) f
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f
hfiber f y ≃ hfp (λ_ : unit, y) f
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y
hfiber f y ≃ hfp (λ_ : unit, y) f
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y
∏ xe : hfiber f y, gg (ff xe) = xe
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y xe: hfiber f y
gg (ff xe) = xe
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y pr1: X pr2: f pr1 = y
gg (ff (pr1,, pr2)) = pr1,, pr2
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe
∏ hf : hfp (λ_ : unit, y) f, ff (gg hf) = hf
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe efg: ∏ hf : hfp (λ_ : unit, y) f, ff (gg hf) = hf
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe hf: hfp (λ_ : unit, y) f
ff (gg hf) = hf
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe efg: ∏ hf : hfp (λ_ : unit, y) f, ff (gg hf) = hf
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe tx: unit × X e: f (pr2 tx) = y
ff (gg (tx,, e)) = tx,, e
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe efg: ∏ hf : hfp (λ_ : unit, y) f, ff (gg hf) = hf
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe t: unit x: X e: f (pr2 (t,, x)) = y
ff (gg ((t,, x),, e)) = (t,, x),, e
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe efg: ∏ hf : hfp (λ_ : unit, y) f, ff (gg hf) = hf
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe x: X e: f (pr2 (tt,, x)) = y
ff (gg ((tt,, x),, e)) = (tt,, x),, e
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe efg: ∏ hf : hfp (λ_ : unit, y) f, ff (gg hf) = hf
isweq ff
X, Y: UU f: X → Y y: Y ff:= hfibertohfp f y: hfiber f y → hfp (λ_ : unit, y) f gg:= hfptohfiber f y: hfp (λ_ : unit, y) f → hfiber f y egf: ∏ xe : hfiber f y, gg (ff xe) = xe efg: ∏ hf : hfp (λ_ : unit, y) f, ff (gg hf) = hf
isweq ff
apply (isweq_iso _ _ egf efg).Defined.
X, Y, Z: UU f: X → Z g: Y → Z
hfp f g ≃ (∑ x : X, hfiber g (f x))
X, Y, Z: UU f: X → Z g: Y → Z
hfp f g ≃ (∑ x : X, hfiber g (f x))
X, Y, Z: UU f: X → Z g: Y → Z
hfp f g ≃ (∑ x : X, hfiber g (f x))
apply weqtotal2dirprodassoc.Defined.
X, Y, Z: UU f: X → Z g: Y → Z
hfp f g ≃ (∑ y : Y, hfiber f (g y))
X, Y, Z: UU f: X → Z g: Y → Z
hfp f g ≃ (∑ y : Y, hfiber f (g y))
X, Y, Z: UU f: X → Z g: Y → Z
hfp f g ≃ (∑ y : Y, hfiber f (g y))
X, Y, Z: UU f: X → Z g: Y → Z
hfp f g → ∑ y : Y, hfiber f (g y)
X, Y, Z: UU f: X → Z g: Y → Z
(∑ y : Y, hfiber f (g y)) → hfp f g
X, Y, Z: UU f: X → Z g: Y → Z
∏ x : hfp f g, ?g (?f x) = x
X, Y, Z: UU f: X → Z g: Y → Z
∏ y : ∑ y : Y, hfiber f (g y), ?f (?g y) = y
X, Y, Z: UU f: X → Z g: Y → Z
hfp f g → ∑ y : Y, hfiber f (g y)
X, Y, Z: UU f: X → Z g: Y → Z x: X y: Y e: g (pr2 (x,, y)) = f (pr1 (x,, y))
∑ y : Y, hfiber f (g y)
exact (y,,x,,!e).
X, Y, Z: UU f: X → Z g: Y → Z
(∑ y : Y, hfiber f (g y)) → hfp f g
X, Y, Z: UU f: X → Z g: Y → Z x: Y y: X e: f y = g x
hfp f g
exact ((y,,x),,!e).
X, Y, Z: UU f: X → Z g: Y → Z
∏ x : hfp f g,
(λX0 : ∑ y : Y, hfiber f (g y),
(λ (x0 : Y) (pr2 : hfiber f (g x0)),
(λ (y : X) (e : f y = g x0), (y,, x0),, ! e)
(pr1 pr2) (Preamble.pr2 pr2)) (pr1 X0) (pr2 X0))
((λX0 : hfp f g,
(λpr1 : X × Y,
(λ (x0 : X) (y : Y)
(e : g (pr2 (x0,, y)) =
f (Preamble.pr1 (x0,, y))), y,, x0,, ! e)
(Preamble.pr1 pr1) (pr2 pr1)) (pr1 X0) (pr2 X0))
x) = x
X, Y, Z: UU f: X → Z g: Y → Z x: X y: Y e: g (pr2 (x,, y)) = f (pr1 (x,, y))
(x,, y),, ! (! e) = (x,, y),, e
apply maponpaths, pathsinv0inv0.
X, Y, Z: UU f: X → Z g: Y → Z
∏ y : ∑ y : Y, hfiber f (g y),
(λX0 : hfp f g,
(λpr1 : X × Y,
(λ (x : X) (y0 : Y)
(e : g (pr2 (x,, y0)) = f (Preamble.pr1 (x,, y0))),
y0,, x,, ! e) (Preamble.pr1 pr1) (pr2 pr1))
(pr1 X0) (pr2 X0))
((λX0 : ∑ y0 : Y, hfiber f (g y0),
(λ (x : Y) (pr2 : hfiber f (g x)),
(λ (y0 : X) (e : f y0 = g x), (y0,, x),, ! e)
(pr1 pr2) (Preamble.pr2 pr2)) (pr1 X0) (pr2 X0))
y) = y
X, Y, Z: UU f: X → Z g: Y → Z x: Y y: X e: f y = g x
(∑ x : X, hfiber g (f x)) ≃ (∑ y : Y, hfiber f (g y))
X, Y, Z: UU f: X → Z g: Y → Z
(∑ x : X, hfiber g (f x)) ≃ (∑ y : Y, hfiber f (g y))
X, Y, Z: UU f: X → Z g: Y → Z
(∑ x : X, hfiber g (f x)) ≃ (∑ y : Y, hfiber f (g y))
X, Y, Z: UU f: X → Z g: Y → Z
(∑ x : X, hfiber g (f x)) → ∑ y : Y, hfiber f (g y)
X, Y, Z: UU f: X → Z g: Y → Z
(∑ y : Y, hfiber f (g y)) → ∑ x : X, hfiber g (f x)
X, Y, Z: UU f: X → Z g: Y → Z
∏ x : ∑ x : X, hfiber g (f x), ?g (?f x) = x
X, Y, Z: UU f: X → Z g: Y → Z
∏ y : ∑ y : Y, hfiber f (g y), ?f (?g y) = y
X, Y, Z: UU f: X → Z g: Y → Z
(∑ x : X, hfiber g (f x)) → ∑ y : Y, hfiber f (g y)
X, Y, Z: UU f: X → Z g: Y → Z x: X y: Y e: g y = f x
∑ y : Y, hfiber f (g y)
exact (y,,x,,!e).
X, Y, Z: UU f: X → Z g: Y → Z
(∑ y : Y, hfiber f (g y)) → ∑ x : X, hfiber g (f x)
X, Y, Z: UU f: X → Z g: Y → Z y: Y x: X e: f x = g y
∑ x : X, hfiber g (f x)
exact (x,,y,,!e).
X, Y, Z: UU f: X → Z g: Y → Z
∏ x : ∑ x : X, hfiber g (f x),
(λX0 : ∑ y : Y, hfiber f (g y),
(λ (y : Y) (pr2 : hfiber f (g y)),
(λ (x0 : X) (e : f x0 = g y), x0,, y,, ! e)
(pr1 pr2) (Preamble.pr2 pr2)) (pr1 X0) (pr2 X0))
((λX0 : ∑ x0 : X, hfiber g (f x0),
(λ (x0 : X) (pr2 : hfiber g (f x0)),
(λ (y : Y) (e : g y = f x0), y,, x0,, ! e)
(pr1 pr2) (Preamble.pr2 pr2)) (pr1 X0) (pr2 X0))
x) = x
X, Y, Z: UU f: X → Z g: Y → Z x: X y: Y e: g y = f x
x,, y,, ! (! e) = x,, y,, e
apply maponpaths, maponpaths, pathsinv0inv0.
X, Y, Z: UU f: X → Z g: Y → Z
∏ y : ∑ y : Y, hfiber f (g y),
(λX0 : ∑ x : X, hfiber g (f x),
(λ (x : X) (pr2 : hfiber g (f x)),
(λ (y0 : Y) (e : g y0 = f x), y0,, x,, ! e)
(pr1 pr2) (Preamble.pr2 pr2)) (pr1 X0) (pr2 X0))
((λX0 : ∑ y0 : Y, hfiber f (g y0),
(λ (y0 : Y) (pr2 : hfiber f (g y0)),
(λ (x : X) (e : f x = g y0), x,, y0,, ! e)
(pr1 pr2) (Preamble.pr2 pr2)) (pr1 X0) (pr2 X0))
y) = y
X, Y, Z: UU f: X → Z g: Y → Z y: Y x: X e: f x = g y
y,, x,, ! (! e) = y,, x,, e
apply maponpaths, maponpaths, pathsinv0inv0.Defined.(** *** Homotopy fiber squares *)Definitionishfsq {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X') (h : commsqstr g' f' g f)
:= isweq (commsqZtohfp f f' g g' h).Definitionhfsqstr {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X')
:= total2 (λh : commsqstr g' f' g f, isweq (commsqZtohfp f f' g g' h)).Definitionmake_hfsqstr {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X')
:= tpair (λh : commsqstr g' f' g f, isweq (commsqZtohfp f f' g g' h)).Definitionhfsqstrtocommsqstr {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X') :
hfsqstr f f' g g' -> commsqstr g' f' g f
:= @pr1 _ (λh : commsqstr g' f' g f, isweq (commsqZtohfp f f' g g' h)).Coercionhfsqstrtocommsqstr : hfsqstr >-> commsqstr.DefinitionweqZtohfp {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y) (g : Z -> X)
(g' : Z -> X') (hf : hfsqstr f f' g g') : Z ≃ hfp f f'
:= make_weq _ (pr2 hf).
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f'
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f'
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y
∏ z : Z, d (c z) = b0 (a z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
d (c z) = b0 (a z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z) =
b0 (tococonusf g z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h) (tococonusf g z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h) (tococonusf g z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h) (tococonusf g z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h)
(g z,, make_hfiber g z (idpath (g z)))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h)
(g z,, make_hfiber g z (idpath (g z)))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
g z,, g' z,, pr1 hf z =
g z,, g' z,, h z @ idpath (f (g z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
h z = h z @ idpath (f (g z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z e: h z = h z @ idpath (f (g z))
g z,, g' z,, pr1 hf z =
g z,, g' z,, h z @ idpath (f (g z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z e: h z = h z @ idpath (f (g z))
g z,, g' z,, pr1 hf z =
g z,, g' z,, h z @ idpath (f (g z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y z: Z
g z,, g' z,, pr1 hf z = g z,, g' z,, h z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (λz : Z, b0 (a z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z) is1: isweq (λz : Z, b0 (a z))
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z)
isweq (λz : Z, d (c z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z) is1: isweq (λz : Z, b0 (a z))
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z) is1: isweq (λz : Z, b0 (a z))
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z) is1: isweq (λz : Z, b0 (a z))
isweq b0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z) is1: isweq (λz : Z, b0 (a z)) is2: isweq b0
isweq (hfibersgtof' f f' g g' hf x)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x: X is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a:= weqtococonusf g: Z ≃ coconusf g c:= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b0:= totalfun (hfiber g) (λx : X, hfiber f' (f x))
(hfibersgtof' f f' g g' h): (∑ y : X, hfiber g y)
→ ∑ y : X, (λx : X, hfiber f' (f x)) y h1: ∏ z : Z, d (c z) = b0 (a z) is1: isweq (λz : Z, b0 (a z)) is2: isweq b0
isweq (hfibersgtof' f f' g g' hf x)
apply (isweqtotaltofib _ _ _ is2 x).Defined.Definitionweqhfibersgtof' {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X') (hf : hfsqstr f f' g g') (x : X)
:= make_weq _ (isweqhfibersgtof' _ _ _ _ hf x).
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x)
hfsqstr f f' g g'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x)
hfsqstr f f' g g'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x)
hfsqstr f f' g g'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x)
isweq (commsqZtohfp f f' g g' h)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g
isweq (commsqZtohfp f f' g g' h)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f'
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f'
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x)
∏ z : Z, d (c0 z) = b (a z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
d (c0 z) = b (a z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(c0 z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(λx : X, hfibersgtof' f f' g g' h x)
(tococonusf g z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(c0 z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(λx : X, hfibersgtof' f f' g g' h x)
(tococonusf g z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(c0 z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(λx : X, hfibersgtof' f f' g g' h x)
(tococonusf g z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(c0 z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(λx : X, hfibersgtof' f f' g g' h x)
(tococonusf g z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(c0 z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(λx : X, hfibersgtof' f f' g g' h x)
(g z,, make_hfiber g z (idpath (g z)))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
total2asstor (λ_ : X, X')
(λxx' : ∑ _ : X, X', f' (pr2 xx') = f (pr1 xx'))
(c0 z) =
totalfun (hfiber g) (λx : X, hfiber f' (f x))
(λx : X, hfibersgtof' f f' g g' h x)
(g z,, make_hfiber g z (idpath (g z)))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
g z,, g' z,, h z = g z,, g' z,, h z @ idpath (f (g z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
h z = h z @ idpath (f (g z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z e: h z = h z @ idpath (f (g z))
g z,, g' z,, h z = g z,, g' z,, h z @ idpath (f (g z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z e: h z = h z @ idpath (f (g z))
g z,, g' z,, h z = g z,, g' z,, h z @ idpath (f (g z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) z: Z
g z,, g' z,, h z = g z,, g' z,, h z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq (λz : Z, d (c0 z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z) is1: isweq (λz : Z, d (c0 z))
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z)
isweq (λz : Z, b (a z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z) is1: isweq (λz : Z, d (c0 z))
isweq c0
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x : X, isweq (hfibersgtof' f f' g g' h x) a:= weqtococonusf g: Z ≃ coconusf g c0:= commsqZtohfp f f' g g' h: Z → hfp f f' d:= weqhfptohfpoverX f f': hfp f f' ≃ hfpoverX f f' b:= weqfibtototal (hfiber g) (λx : X, hfiber f' (f x))
(λx : X,
make_weq (hfibersgtof' f f' g g' h x) (is x)): (∑ x : X, hfiber g x)
≃ (∑ x : X, (λx0 : X, hfiber f' (f x0)) x) h1: ∏ z : Z, d (c0 z) = b (a z) is1: isweq (λz : Z, d (c0 z))
isweq c0
apply (twooutof3a _ _ is1 (pr2 d)).Defined.
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X'
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X'
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X'
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g'
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f'
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f'
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y
∏ z : Z, d' (c' z) = b0' (a' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z
d' (c' z) = b0' (a' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z
d' (c' z) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h) (a' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z
d' (c' z) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h) (weqtococonusf g' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z
d' (c' z) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h)
(make_weq (tococonusf g') (isweqtococonusf g') z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z
d' (c' z) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h)
(make_weq
(λx : Z, g' x,, make_hfiber g' x (idpath (g' x)))
(isweqtococonusf g') z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z))) =
g' z,, g z,, ! h z @ idpath (f' (g' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z
! h z = ! h z @ idpath (f' (g' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z e: ! h z = ! h z @ idpath (f' (g' z))
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z))) =
g' z,, g z,, ! h z @ idpath (f' (g' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z e: ! h z = ! h z @ idpath (f' (g' z))
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z))) =
g' z,, g z,, ! h z @ idpath (f' (g' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y z: Z
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(commsqZtohfp f f' g g' (pr1 hf) z))) =
g' z,, g z,, ! h z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (λz : Z, b0' (a' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z) is1: isweq (λz : Z, b0' (a' z))
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z)
isweq (λz : Z, d' (c' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z) is1: isweq (λz : Z, b0' (a' z))
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z) is1: isweq (λz : Z, b0' (a' z))
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z) is1: isweq (λz : Z, b0' (a' z))
isweq b0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z) is1: isweq (λz : Z, b0' (a' z)) is2: isweq b0'
isweq (hfibersg'tof f f' g g' hf x')
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' x': X' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f a':= weqtococonusf g': Z ≃ coconusf g' c':= make_weq (commsqZtohfp f f' g g' (pr1 hf)) is: Z ≃ hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b0':= totalfun (hfiber g')
(λx' : X', hfiber f (f' x'))
(hfibersg'tof f f' g g' h): (∑ y : X', hfiber g' y) → ∑ y : X', (λx' : X', hfiber f (f' x')) y h1: ∏ z : Z, d' (c' z) = b0' (a' z) is1: isweq (λz : Z, b0' (a' z)) is2: isweq b0'
isweq (hfibersg'tof f f' g g' hf x')
apply (isweqtotaltofib _ _ _ is2 x').Defined.Definitionweqhfibersg'tof {XX'YZ : UU} (f : X -> Y) (f' : X' -> Y)
(g : Z -> X) (g' : Z -> X') (hf : hfsqstr f f' g g') (x' : X')
:= make_weq _ (isweqhfibersg'tof _ _ _ _ hf x').
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x')
hfsqstr f f' g g'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x')
hfsqstr f f' g g'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x')
hfsqstr f f' g g'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x')
isweq (commsqZtohfp f f' g g' h)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g'
isweq (commsqZtohfp f f' g g' h)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f'
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f'
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x)
∏ z : Z, d' (c0' z) = b' (a' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
d' (c0' z) = b' (a' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(c0' z))) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(λx : X', hfibersg'tof f f' g g' h x)
(tococonusf g' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(c0' z))) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(λx : X', hfibersg'tof f f' g g' h x)
(tococonusf g' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(c0' z))) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(λx : X', hfibersg'tof f f' g g' h x)
(tococonusf g' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(c0' z))) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(λx : X', hfibersg'tof f f' g g' h x)
(tococonusf g' z)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
total2asstor (λ_ : X', X)
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(totalfun
(λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(c0' z))) =
totalfun (hfiber g') (λx' : X', hfiber f (f' x'))
(λx : X', hfibersg'tof f f' g g' h x)
(g' z,, make_hfiber g' z (idpath (g' z)))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
g' z,, g z,, ! h z =
g' z,, g z,, ! h z @ idpath (f' (g' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
! h z = ! h z @ idpath (f' (g' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z e: ! h z = ! h z @ idpath (f' (g' z))
g' z,, g z,, ! h z =
g' z,, g z,, ! h z @ idpath (f' (g' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z e: ! h z = ! h z @ idpath (f' (g' z))
g' z,, g z,, ! h z =
g' z,, g z,, ! h z @ idpath (f' (g' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) z: Z
g' z,, g z,, ! h z = g' z,, g z,, ! h z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq (λz : Z, d' (c0' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z) is1: isweq (λz : Z, d' (c0' z))
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z)
isweq (λz : Z, b' (a' z))
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z) is1: isweq (λz : Z, d' (c0' z))
isweq c0'
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' h: commsqstr g' f' g f is: ∏ x' : X', isweq (hfibersg'tof f f' g g' h x') a':= weqtococonusf g': Z ≃ coconusf g' c0':= commsqZtohfp f f' g g' h: Z → hfp f f' d':= weqhfptohfpoverX' f f': hfp f f' ≃ hfpoverX' f f' b':= weqfibtototal (hfiber g')
(λx' : X', hfiber f (f' x'))
(λx' : X',
make_weq (hfibersg'tof f f' g g' h x') (is x')): (∑ x : X', hfiber g' x)
≃ (∑ x : X', (λx' : X', hfiber f (f' x')) x) h1: ∏ z : Z, d' (c0' z) = b' (a' z) is1: isweq (λz : Z, d' (c0' z))
isweq c0'
apply (twooutof3a _ _ is1 (pr2 d')).Defined.
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g'
hfsqstr f' f g' g
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g'
hfsqstr f' f g' g
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g'
hfsqstr f' f g' g
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf)
hfsqstr f' f g' g
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f
hfsqstr f' f g' g
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f'
hfsqstr f' f g' g
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f'
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f
∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f z: Z
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f z: Z
w1 (make_dirprod (g z) (g' z),, h z) =
make_dirprod (g' z) (g z),, th z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f z: Z
totalfun (λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(make_dirprod (g z) (g' z),, h z)) =
make_dirprod (g' z) (g z),, th z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f z: Z
totalfun (λx'x : X' × X, f' (pr1 x'x) = f (pr2 x'x))
(λx'x : X' × X, f (pr2 x'x) = f' (pr1 x'x))
(λx : X' × X, pathsinv0)
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(make_dirprod (g z) (g' z),, h z)) =
make_dirprod (g' z) (g z),, th z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f z: Z
pr1
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(make_dirprod (g z) (g' z),, h z)),,
! pr2
(weqfp_map (weqdirprodcomm X X')
(λxx' : X' × X, f' (pr1 xx') = f (pr2 xx'))
(make_dirprod (g z) (g' z),, h z)) =
make_dirprod (g' z) (g z),, th z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f z: Z
make_dirprod (g' z) (g z),, ! h z =
make_dirprod (g' z) (g z),, th z
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (commsqZtohfp f' f g' g th)
X, X', Y, Z: UU f: X → Y f': X' → Y g: Z → X g': Z → X' hf: hfsqstr f f' g g' is:= pr2 hf: (λh : commsqstr g' f' g f,
isweq (commsqZtohfp f f' g g' h)) (pr1 hf) h:= pr1 hf: commsqstr g' f' g f th:= transposcommsqstr f f' g g' h: commsqstr g f g' f' w1:= weqhfpcomm f f': hfp f f' ≃ hfp f' f h1: ∏ z : Z,
w1 (commsqZtohfp f f' g g' h z) =
commsqZtohfp f' f g' g th z
isweq (λz : Z, w1 (commsqZtohfp f f' g g' h z))
apply (twooutof3c _ _ is (pr2 w1)).Defined.(** *** Fiber sequences and homotopy fiber squares *)
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: fibseqstr f g z
hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: fibseqstr f g z
hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: fibseqstr f g z
hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: fibseqstr f g z
isweq
(commsqZtohfp (λ_ : unit, z) g (λ_ : X, tt) f
(pr1 hf))
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: fibseqstr f g z ff:= ezweq f g z hf: X ≃ hfiber g z
isweq
(commsqZtohfp (λ_ : unit, z) g (λ_ : X, tt) f
(pr1 hf))
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: fibseqstr f g z ff:= ezweq f g z hf: X ≃ hfiber g z ggff:= commsqZtohfp (λ_ : unit, z) g (λ_ : X, tt) f (pr1 hf): X → hfp (λ_ : unit, z) g
isweq ggff
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: fibseqstr f g z ff:= ezweq f g z hf: X ≃ hfiber g z ggff:= commsqZtohfp (λ_ : unit, z) g
(λ_ : X, tt) f (pr1 hf): X → hfp (λ_ : unit, z) g gg:= weqhfibertohfp g z: hfiber g z ≃ hfp (λ_ : unit, z) g
isweq ggff
apply (pr2 (weqcomp ff gg)).Defined.
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f
fibseqstr f g z
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f
fibseqstr f g z
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f
fibseqstr f g z
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f
isfibseq f g z (pr1 hf)
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f ff:= ezmap f g z (pr1 hf): X → hfiber g z
isfibseq f g z (pr1 hf)
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f ff:= ezmap f g z (pr1 hf): X → hfiber g z ggff:= weqZtohfp (λ_ : unit, z) g (λ_ : X, tt) f hf: X ≃ hfp (λ_ : unit, z) g
isfibseq f g z (pr1 hf)
X, Y, Z: UU f: X → Y g: Y → Z z: Z hf: hfsqstr (λ_ : unit, z) g (λ_ : X, tt) f ff:= ezmap f g z (pr1 hf): X → hfiber g z ggff:= weqZtohfp (λ_ : unit, z) g (λ_ : X, tt) f hf: X ≃ hfp (λ_ : unit, z) g gg:= weqhfibertohfp g z: hfiber g z ≃ hfp (λ_ : unit, z) g
isfibseq f g z (pr1 hf)
apply (twooutof3a ff gg (pr2 ggff) (pr2 gg)).Defined.(* End of the file PartA.v *)